Library prosa.analysis.abstract.ideal_jlfp_rta


(* ----------------------------------[ coqtop ]---------------------------------

Welcome to Coq 8.11.2 (June 2020)

----------------------------------------------------------------------------- *)


Require Export prosa.analysis.definitions.priority_inversion.
Require Export prosa.analysis.abstract.abstract_seq_rta.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq path fintype bigop.

In this file we consider an ideal uni-processor ...
Require Import prosa.model.processor.ideal.

... and classic model of readiness without jitter and no self-suspensions, where pending jobs are always ready.
Require Import prosa.model.readiness.basic.

JLFP instantiation of Interference and Interfering Workload for ideal uni-processor.

In this module we instantiate functions Interference and Interfering Workload for an arbitrary JLFP-policy that satisfies the sequential tasks hypothesis. We also prove equivalence of Interference and Interfering Workload to the more conventional notions of service or workload.
Consider any type of tasks ...
  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 arrival sequence with consistent arrivals.
Next, consider any ideal uni-processor schedule of this arrival sequence ...
... where jobs do not execute before their arrival or after completion.
Assume we have sequential tasks, i.e., jobs of the same task execute in the order of their arrival.
Consider a JLFP-policy that indicates a higher-or-equal priority relation, and assume that this relation is reflexive and transitive.
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.
Let [tsk] be any task in ts that is to be analyzed.
  Variable tsk : Task.

For simplicity, let's define some local names.

Interference and Interfering Workload

In this section, we introduce definitions of interference, interfering workload and a function that bounds cumulative interference.
For proper calculation of interference and interfering workload of a job, we need to distinguish interference received from other jobs of the same task and jobs of other tasks. In that regard, we introduce two additional relations. The first relation defines whether job [j1] has a higher-than-or-equal-priority than job [j2] and [j1] is not equal to [j2]...
  Let another_hep_job: JLFP_policy Job :=
    fun j1 j2hep_job j1 j2 && (j1 != j2).

...and the second relation defines whether a job [j1] has a higher-or-equal-priority than job [j2] and the task of [j1] is not equal to task of [j2].
In order to introduce the interference, first we need to recall the definition of priority inversion introduced in module limited.fixed_priority.busy_interval: [ Definition is_priority_inversion t := ] [ if sched t is Some jlp then ] [ ~~ higher_eq_priority jlp j ] [ else false. ] I.e., 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.
  Let is_priority_inversion (j : Job) (t : instant) :=
    is_priority_inversion sched j t.

Next, we say that job j is incurring interference from another job with higher or equal priority at time t, if there exists job [jhp] (different from j) with a higher or equal priority that executes at time t.
  Definition is_interference_from_another_hep_job (j : Job) (t : instant) :=
    if sched t is Some jhp then
      another_hep_job jhp j
    else false.

Similarly, we say that job j is incurring interference from a job with higher or equal priority of another task at time t, if there exists a job [jhp] (of a different task) with higher or equal priority that executes at time t.
  Definition is_interference_from_hep_job_from_another_task (j : Job) (t : instant) :=
    if sched t is Some jhp then
      hep_job_from_another_task jhp j
    else false.

Now, we define the notion of cumulative interference, called interfering_workload_of_jobs_with_hep_priority, that says how many units of workload are generated by jobs with higher or equal priority released at time t.
Instantiation of Interference 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.
Instantiation of Interfering Workload The interfering workload, in turn, is defined as the sum of the priority inversion function and interfering workload of jobs with higher or equal priority.
For each of the concepts defined above, we introduce a corresponding cumulative function: (a) cumulative priority inversion...
  Definition cumulative_priority_inversion (j : Job) (t1 t2 : instant) :=
    \sum_(t1 t < t2) is_priority_inversion j t.

... (b) cumulative interference from other jobs with higher or equal priority...
... (c) and cumulative interference from jobs with higher or equal priority from other tasks...
... (d) cumulative interference...
... (e) cumulative workload from jobs with higher or equal priority...
... (f) and cumulative interfering workload.
Instantiated functions usually do not have any useful lemmas about them. In order to reuse existing lemmas, we need to prove equivalence of the instantiated functions to some conventional notions. The instantiations given in this file are equivalent to service and workload. Further, we prove these equivalences formally.
Before we present the formal proofs of the equivalences, we recall the notion of workload of higher or equal priority jobs.
  Let workload_of_other_hep_jobs (j : Job) (t1 t2 : instant) :=
    workload_of_jobs (fun jhpanother_hep_job jhp j) (arrivals_between t1 t2).

Similarly, we recall notions of service of higher or equal priority jobs from other tasks...
  Let service_of_hep_jobs_from_other_tasks (j : Job) (t1 t2 : instant) :=
    service_of_jobs sched (fun jhphep_job_from_another_task jhp j)
                    (arrivals_between t1 t2) t1 t2.

... and service of all other jobs with higher or equal priority.
  Let service_of_other_hep_jobs (j : Job) (t1 t2 : instant) :=
    service_of_jobs sched (fun jhpanother_hep_job jhp j) (arrivals_between t1 t2) t1 t2.

Equivalences

In this section we prove a few equivalences between the definitions obtained by instantiation of definitions from the Abstract RTA module (interference and interfering workload) and definitions corresponding to the conventional concepts.
As it was mentioned previously, instantiated functions of interference and interfering workload usually do not have any useful lemmas about them. However, it is possible to prove their equivalence to the more conventional notions like service or workload. Next we prove the equivalence between the instantiations and conventional notions.
  Section Equivalences.

We prove that we can split cumulative interference into two parts: (1) cumulative priority inversion and (2) cumulative interference from jobs with higher or equal priority.
    Lemma cumulative_interference_split:
       j t1 t2,
        cumulative_interference j t1 t2
        = cumulative_priority_inversion j t1 t2
          + cumulative_interference_from_other_hep_jobs j t1 t2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1301)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  ============================
  forall (j : Job) (t1 t2 : instant),
  cumulative_interference j t1 t2 =
  cumulative_priority_inversion j t1 t2 +
  cumulative_interference_from_other_hep_jobs j t1 t2

----------------------------------------------------------------------------- *)


    Proof.
      rewrite /cumulative_interference /interference.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1303)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  ============================
  forall (j : Job) (t1 t2 : instant),
  \sum_(t1 <= t < t2)
     (is_priority_inversion j t || is_interference_from_another_hep_job j t) =
  cumulative_priority_inversion j t1 t2 +
  cumulative_interference_from_other_hep_jobs j t1 t2

----------------------------------------------------------------------------- *)


      intros; rewrite -big_split //=.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1317)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1, t2 : instant
  ============================
  \sum_(t1 <= t < t2)
     (is_priority_inversion j t || is_interference_from_another_hep_job j t) =
  \sum_(t1 <= i < t2)
     (is_priority_inversion j i + is_interference_from_another_hep_job j i)

----------------------------------------------------------------------------- *)


      apply/eqP; rewrite eqn_leq; apply/andP; split; rewrite leq_sum; try done.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1434)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1, t2 : instant
  ============================
  forall i : nat,
  true ->
  is_priority_inversion j i || is_interference_from_another_hep_job j i <=
  is_priority_inversion j i + is_interference_from_another_hep_job j i

subgoal 2 (ID 1443) is:
 forall i : nat,
 true ->
 is_priority_inversion j i + is_interference_from_another_hep_job j i <=
 is_priority_inversion j i || is_interference_from_another_hep_job j i

----------------------------------------------------------------------------- *)


      {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1434)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1, t2 : instant
  ============================
  forall i : nat,
  true ->
  is_priority_inversion j i || is_interference_from_another_hep_job j i <=
  is_priority_inversion j i + is_interference_from_another_hep_job j i

----------------------------------------------------------------------------- *)


intros t _; unfold is_priority_inversion, priority_inversion.is_priority_inversion.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1500)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ============================
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end || is_interference_from_another_hep_job j t <=
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end + is_interference_from_another_hep_job j t

----------------------------------------------------------------------------- *)


        ideal_proc_model_sched_case_analysis_eq sched t s; first by done.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1596)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1, t2 : instant
  t : nat
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  ============================
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end || is_interference_from_another_hep_job j t <=
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end + is_interference_from_another_hep_job j t

----------------------------------------------------------------------------- *)


        destruct (hep_job s j) eqn:MM; simpl; rewrite ?addn0 ?add0n.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1612)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1, t2 : instant
  t : nat
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  MM : hep_job s j = true
  ============================
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end || is_interference_from_another_hep_job j t <=
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end + is_interference_from_another_hep_job j t

subgoal 2 (ID 1614) is:
 match sched t with
 | Some jlp => ~~ hep_job jlp j
 | None => false
 end || is_interference_from_another_hep_job j t <=
 match sched t with
 | Some jlp => ~~ hep_job jlp j
 | None => false
 end + is_interference_from_another_hep_job j t

----------------------------------------------------------------------------- *)


        all: by move: Sched_s; rewrite scheduled_at_def; move ⇒ /eqP EQ; rewrite EQ MM.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1443)

subgoal 1 (ID 1443) is:
 forall i : nat,
 true ->
 is_priority_inversion j i + is_interference_from_another_hep_job j i <=
 is_priority_inversion j i || is_interference_from_another_hep_job j i

----------------------------------------------------------------------------- *)


      }
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1443)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1, t2 : instant
  ============================
  forall i : nat,
  true ->
  is_priority_inversion j i + is_interference_from_another_hep_job j i <=
  is_priority_inversion j i || is_interference_from_another_hep_job j i

----------------------------------------------------------------------------- *)


             
      {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1443)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1, t2 : instant
  ============================
  forall i : nat,
  true ->
  is_priority_inversion j i + is_interference_from_another_hep_job j i <=
  is_priority_inversion j i || is_interference_from_another_hep_job j i

----------------------------------------------------------------------------- *)


intros t _; unfold is_priority_inversion, priority_inversion.is_priority_inversion,
                    is_interference_from_another_hep_job.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1725)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ============================
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end +
  match sched t with
  | Some jhp => another_hep_job jhp j
  | None => false
  end <=
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end
  || match sched t with
     | Some jhp => another_hep_job jhp j
     | None => false
     end

----------------------------------------------------------------------------- *)


        ideal_proc_model_sched_case_analysis_eq sched t s; first by done.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1821)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1, t2 : instant
  t : nat
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  ============================
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end +
  match sched t with
  | Some jhp => another_hep_job jhp j
  | None => false
  end <=
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end
  || match sched t with
     | Some jhp => another_hep_job jhp j
     | None => false
     end

----------------------------------------------------------------------------- *)


        unfold another_hep_job.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1824)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1, t2 : instant
  t : nat
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  ============================
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end +
  match sched t with
  | Some jhp => hep_job jhp j && (jhp != j)
  | None => false
  end <=
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end
  || match sched t with
     | Some jhp => hep_job jhp j && (jhp != j)
     | None => false
     end

----------------------------------------------------------------------------- *)


        destruct (hep_job s j) eqn:HP; simpl; rewrite ?addn0 ?add0n.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1839)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1, t2 : instant
  t : nat
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  HP : hep_job s j = true
  ============================
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end +
  match sched t with
  | Some jhp => hep_job jhp j && (jhp != j)
  | None => false
  end <=
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end
  || match sched t with
     | Some jhp => hep_job jhp j && (jhp != j)
     | None => false
     end

subgoal 2 (ID 1841) is:
 match sched t with
 | Some jlp => ~~ hep_job jlp j
 | None => false
 end +
 match sched t with
 | Some jhp => hep_job jhp j && (jhp != j)
 | None => false
 end <=
 match sched t with
 | Some jlp => ~~ hep_job jlp j
 | None => false
 end
 || match sched t with
    | Some jhp => hep_job jhp j && (jhp != j)
    | None => false
    end

----------------------------------------------------------------------------- *)


        all: by move: Sched_s; rewrite scheduled_at_def; move ⇒ /eqP EQ; rewrite EQ HP.

(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


      }

(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


    Qed.

Let [j] be any job of task [tsk], and let [upp_t] be any time instant after job [j]'s arrival. Then for any time interval lying before [upp_t], the cumulative interference received by [tsk] is equal to the sum of the cumulative priority inversion of job [j] and the cumulative interference incurred by task [tsk] due to other tasks.
    Lemma cumulative_task_interference_split:
       j t1 t2 upp_t,
        arrives_in arr_seq j
        job_task j = tsk
        j \in arrivals_before arr_seq upp_t
        ~~ job_completed_by j t2
        cumulative_task_interference interference tsk upp_t t1 t2 =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_hep_jobs_from_other_tasks j t1 t2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1319)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  ============================
  forall (j : Job) (t1 : nat) (t2 upp_t : instant),
  arrives_in arr_seq j ->
  job_task j = tsk ->
  j \in arrivals_before arr_seq upp_t ->
  ~~ job_completed_by j t2 ->
  cumulative_task_interference interference tsk upp_t t1 t2 =
  cumulative_priority_inversion j t1 t2 +
  cumulative_interference_from_hep_jobs_from_other_tasks j t1 t2

----------------------------------------------------------------------------- *)


    Proof.
      rewrite /cumulative_task_interference /cumul_task_interference.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1327)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  ============================
  forall (j : Job) (t1 : nat) (t2 upp_t : instant),
  arrives_in arr_seq j ->
  job_task j = tsk ->
  j \in arrivals_before arr_seq upp_t ->
  ~~ job_completed_by j t2 ->
  \sum_(t1 <= t < t2)
     task_interference_received_before arr_seq sched interference tsk upp_t t =
  cumulative_priority_inversion j t1 t2 +
  cumulative_interference_from_hep_jobs_from_other_tasks j t1 t2

----------------------------------------------------------------------------- *)


      intros j t1 R upp ARRin TSK ARR NCOMPL.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1335)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  ============================
  \sum_(t1 <= t < R)
     task_interference_received_before arr_seq sched interference tsk upp t =
  cumulative_priority_inversion j t1 R +
  cumulative_interference_from_hep_jobs_from_other_tasks j t1 R

----------------------------------------------------------------------------- *)


      rewrite -big_split //= big_nat_cond [X in _ = X]big_nat_cond.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1394)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  ============================
  \sum_(t1 <= i < R | (t1 <= i < R) && true)
     task_interference_received_before arr_seq sched interference tsk upp i =
  \sum_(t1 <= i < R | (t1 <= i < R) && true)
     (is_priority_inversion j i +
      is_interference_from_hep_job_from_another_task j i)

----------------------------------------------------------------------------- *)


      apply/eqP; rewrite eqn_leq; apply/andP; split.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1479)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  ============================
  \sum_(t1 <= i < R | (t1 <= i < R) && true)
     task_interference_received_before arr_seq sched interference tsk upp i <=
  \sum_(t1 <= i < R | (t1 <= i < R) && true)
     (is_priority_inversion j i +
      is_interference_from_hep_job_from_another_task j i)

subgoal 2 (ID 1480) is:
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    (is_priority_inversion j i +
     is_interference_from_hep_job_from_another_task j i) <=
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    task_interference_received_before arr_seq sched interference tsk upp i

----------------------------------------------------------------------------- *)


      all: rewrite /interference /is_priority_inversion /priority_inversion.is_priority_inversion.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1485)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  ============================
  \sum_(t1 <= i < R | (t1 <= i < R) && true)
     task_interference_received_before arr_seq sched
       (fun (j0 : Job) (t : instant) =>
        match sched t with
        | Some jlp => ~~ hep_job jlp j0
        | None => false
        end || is_interference_from_another_hep_job j0 t) tsk upp i <=
  \sum_(t1 <= i < R | (t1 <= i < R) && true)
     (match sched i with
      | Some jlp => ~~ hep_job jlp j
      | None => false
      end + is_interference_from_hep_job_from_another_task j i)

subgoal 2 (ID 1490) is:
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    (match sched i with
     | Some jlp => ~~ hep_job jlp j
     | None => false
     end + is_interference_from_hep_job_from_another_task j i) <=
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    task_interference_received_before arr_seq sched
      (fun (j0 : Job) (t : instant) =>
       match sched t with
       | Some jlp => ~~ hep_job jlp j0
       | None => false
       end || is_interference_from_another_hep_job j0 t) tsk upp i

----------------------------------------------------------------------------- *)


      - apply leq_sum; intros t _.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1494)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  ============================
  task_interference_received_before arr_seq sched
    (fun (j0 : Job) (t0 : instant) =>
     match sched t0 with
     | Some jlp => ~~ hep_job jlp j0
     | None => false
     end || is_interference_from_another_hep_job j0 t0) tsk upp t <=
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end + is_interference_from_hep_job_from_another_task j t

subgoal 2 (ID 1490) is:
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    (match sched i with
     | Some jlp => ~~ hep_job jlp j
     | None => false
     end + is_interference_from_hep_job_from_another_task j i) <=
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    task_interference_received_before arr_seq sched
      (fun (j0 : Job) (t : instant) =>
       match sched t with
       | Some jlp => ~~ hep_job jlp j0
       | None => false
       end || is_interference_from_another_hep_job j0 t) tsk upp i

----------------------------------------------------------------------------- *)


        rewrite /task_interference_received_before /task_schedule.task_scheduled_at
                /is_interference_from_hep_job_from_another_task
                /is_interference_from_another_hep_job /hep_job_from_another_task.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1511)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  ============================
  ~~ match sched t with
     | Some j0 => job_task j0 == tsk
     | None => false
     end &&
  has
    (fun j0 : Job =>
     match sched t with
     | Some jlp => ~~ hep_job jlp j0
     | None => false
     end
     || match sched t with
        | Some jhp => another_hep_job jhp j0
        | None => false
        end) (task_arrivals_before arr_seq tsk upp) <=
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end +
  match sched t with
  | Some jhp => hep_job jhp j && (job_task jhp != job_task j)
  | None => false
  end

subgoal 2 (ID 1490) is:
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    (match sched i with
     | Some jlp => ~~ hep_job jlp j
     | None => false
     end + is_interference_from_hep_job_from_another_task j i) <=
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    task_interference_received_before arr_seq sched
      (fun (j0 : Job) (t : instant) =>
       match sched t with
       | Some jlp => ~~ hep_job jlp j0
       | None => false
       end || is_interference_from_another_hep_job j0 t) tsk upp i

----------------------------------------------------------------------------- *)


        ideal_proc_model_sched_case_analysis_eq sched t s; first by rewrite has_pred0 addn0 leqn0 eqb0.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1607)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  ============================
  ~~ match sched t with
     | Some j0 => job_task j0 == tsk
     | None => false
     end &&
  has
    (fun j0 : Job =>
     match sched t with
     | Some jlp => ~~ hep_job jlp j0
     | None => false
     end
     || match sched t with
        | Some jhp => another_hep_job jhp j0
        | None => false
        end) (task_arrivals_before arr_seq tsk upp) <=
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end +
  match sched t with
  | Some jhp => hep_job jhp j && (job_task jhp != job_task j)
  | None => false
  end

subgoal 2 (ID 1490) is:
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    (match sched i with
     | Some jlp => ~~ hep_job jlp j
     | None => false
     end + is_interference_from_hep_job_from_another_task j i) <=
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    task_interference_received_before arr_seq sched
      (fun (j0 : Job) (t : instant) =>
       match sched t with
       | Some jlp => ~~ hep_job jlp j0
       | None => false
       end || is_interference_from_another_hep_job j0 t) tsk upp i

----------------------------------------------------------------------------- *)


        destruct (hep_job s j) eqn:HP; simpl.

(* ----------------------------------[ coqtop ]---------------------------------

3 subgoals (ID 1638)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  HP : hep_job s j = true
  ============================
  ~~ match sched t with
     | Some j0 => job_task j0 == tsk
     | None => false
     end &&
  has
    (fun j0 : Job =>
     match sched t with
     | Some jlp => ~~ hep_job jlp j0
     | None => false
     end
     || match sched t with
        | Some jhp => another_hep_job jhp j0
        | None => false
        end) (task_arrivals_before arr_seq tsk upp) <=
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end +
  match sched t with
  | Some jhp => hep_job jhp j && (job_task jhp != job_task j)
  | None => false
  end

subgoal 2 (ID 1640) is:
 ~~ match sched t with
    | Some j0 => job_task j0 == tsk
    | None => false
    end &&
 has
   (fun j0 : Job =>
    match sched t with
    | Some jlp => ~~ hep_job jlp j0
    | None => false
    end
    || match sched t with
       | Some jhp => another_hep_job jhp j0
       | None => false
       end) (task_arrivals_before arr_seq tsk upp) <=
 match sched t with
 | Some jlp => ~~ hep_job jlp j
 | None => false
 end +
 match sched t with
 | Some jhp => hep_job jhp j && (job_task jhp != job_task j)
 | None => false
 end
subgoal 3 (ID 1490) is:
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    (match sched i with
     | Some jlp => ~~ hep_job jlp j
     | None => false
     end + is_interference_from_hep_job_from_another_task j i) <=
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    task_interference_received_before arr_seq sched
      (fun (j0 : Job) (t : instant) =>
       match sched t with
       | Some jlp => ~~ hep_job jlp j0
       | None => false
       end || is_interference_from_another_hep_job j0 t) tsk upp i

----------------------------------------------------------------------------- *)


        1-2: move: Sched_s; rewrite scheduled_at_def; move ⇒ /eqP EQ; rewrite EQ HP.

(* ----------------------------------[ coqtop ]---------------------------------

3 subgoals (ID 1730)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  s : Job
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  HP : hep_job s j = true
  EQ : sched t = Some s
  ============================
  (job_task s != tsk) &&
  has (fun j0 : Job => ~~ hep_job s j0 || another_hep_job s j0)
    (task_arrivals_before arr_seq tsk upp) <=
  ~~ true + true && (job_task s != job_task j)

subgoal 2 (ID 1734) is:
 (job_task s != tsk) &&
 has (fun j0 : Job => ~~ hep_job s j0 || another_hep_job s j0)
   (task_arrivals_before arr_seq tsk upp) <=
 ~~ false + false && (job_task s != job_task j)
subgoal 3 (ID 1490) is:
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    (match sched i with
     | Some jlp => ~~ hep_job jlp j
     | None => false
     end + is_interference_from_hep_job_from_another_task j i) <=
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    task_interference_received_before arr_seq sched
      (fun (j0 : Job) (t : instant) =>
       match sched t with
       | Some jlp => ~~ hep_job jlp j0
       | None => false
       end || is_interference_from_another_hep_job j0 t) tsk upp i

----------------------------------------------------------------------------- *)


        + rewrite add0n TSK.

(* ----------------------------------[ coqtop ]---------------------------------

3 subgoals (ID 1741)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  s : Job
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  HP : hep_job s j = true
  EQ : sched t = Some s
  ============================
  (job_task s != tsk) &&
  has (fun j0 : Job => ~~ hep_job s j0 || another_hep_job s j0)
    (task_arrivals_before arr_seq tsk upp) <= true && (job_task s != tsk)

subgoal 2 (ID 1734) is:
 (job_task s != tsk) &&
 has (fun j0 : Job => ~~ hep_job s j0 || another_hep_job s j0)
   (task_arrivals_before arr_seq tsk upp) <=
 ~~ false + false && (job_task s != job_task j)
subgoal 3 (ID 1490) is:
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    (match sched i with
     | Some jlp => ~~ hep_job jlp j
     | None => false
     end + is_interference_from_hep_job_from_another_task j i) <=
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    task_interference_received_before arr_seq sched
      (fun (j0 : Job) (t : instant) =>
       match sched t with
       | Some jlp => ~~ hep_job jlp j0
       | None => false
       end || is_interference_from_another_hep_job j0 t) tsk upp i

----------------------------------------------------------------------------- *)


          by case: (job_task s != tsk); first rewrite Bool.andb_true_l leq_b1.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1734)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  s : Job
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  HP : hep_job s j = false
  EQ : sched t = Some s
  ============================
  (job_task s != tsk) &&
  has (fun j0 : Job => ~~ hep_job s j0 || another_hep_job s j0)
    (task_arrivals_before arr_seq tsk upp) <=
  ~~ false + false && (job_task s != job_task j)

subgoal 2 (ID 1490) is:
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    (match sched i with
     | Some jlp => ~~ hep_job jlp j
     | None => false
     end + is_interference_from_hep_job_from_another_task j i) <=
 \sum_(t1 <= i < R | (t1 <= i < R) && true)
    task_interference_received_before arr_seq sched
      (fun (j0 : Job) (t : instant) =>
       match sched t with
       | Some jlp => ~~ hep_job jlp j0
       | None => false
       end || is_interference_from_another_hep_job j0 t) tsk upp i

----------------------------------------------------------------------------- *)


        + by rewrite addn0 leq_b1.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1490)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  ============================
  \sum_(t1 <= i < R | (t1 <= i < R) && true)
     (match sched i with
      | Some jlp => ~~ hep_job jlp j
      | None => false
      end + is_interference_from_hep_job_from_another_task j i) <=
  \sum_(t1 <= i < R | (t1 <= i < R) && true)
     task_interference_received_before arr_seq sched
       (fun (j0 : Job) (t : instant) =>
        match sched t with
        | Some jlp => ~~ hep_job jlp j0
        | None => false
        end || is_interference_from_another_hep_job j0 t) tsk upp i

----------------------------------------------------------------------------- *)


      - apply leq_sum; movet /andP [/andP [_ LT'] _].

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1846)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  ============================
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end + is_interference_from_hep_job_from_another_task j t <=
  task_interference_received_before arr_seq sched
    (fun (j0 : Job) (t0 : instant) =>
     match sched t0 with
     | Some jlp => ~~ hep_job jlp j0
     | None => false
     end || is_interference_from_another_hep_job j0 t0) tsk upp t

----------------------------------------------------------------------------- *)


        rewrite /is_interference_from_hep_job_from_another_task
                /hep_job_from_another_task /task_interference_received_before
                /task_schedule.task_scheduled_at.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1862)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  ============================
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end +
  match sched t with
  | Some jhp => hep_job jhp j && (job_task jhp != job_task j)
  | None => false
  end <=
  ~~ match sched t with
     | Some j0 => job_task j0 == tsk
     | None => false
     end &&
  has
    (fun j0 : Job =>
     match sched t with
     | Some jlp => ~~ hep_job jlp j0
     | None => false
     end || is_interference_from_another_hep_job j0 t)
    (task_arrivals_before arr_seq tsk upp)

----------------------------------------------------------------------------- *)


        ideal_proc_model_sched_case_analysis_eq sched t s; first by done.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1958)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  ============================
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end +
  match sched t with
  | Some jhp => hep_job jhp j && (job_task jhp != job_task j)
  | None => false
  end <=
  ~~ match sched t with
     | Some j0 => job_task j0 == tsk
     | None => false
     end &&
  has
    (fun j0 : Job =>
     match sched t with
     | Some jlp => ~~ hep_job jlp j0
     | None => false
     end || is_interference_from_another_hep_job j0 t)
    (task_arrivals_before arr_seq tsk upp)

----------------------------------------------------------------------------- *)


        move: (Sched_s); rewrite scheduled_at_def; move ⇒ /eqP EQ; rewrite EQ.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2003)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  ============================
  ~~ hep_job s j + hep_job s j && (job_task s != job_task j) <=
  (job_task s != tsk) &&
  has
    (fun j0 : Job =>
     ~~ hep_job s j0 || is_interference_from_another_hep_job j0 t)
    (task_arrivals_before arr_seq tsk upp)

----------------------------------------------------------------------------- *)


        rewrite -TSK; case TSKEQ: (job_task s == job_task j); simpl.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2120)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = true
  ============================
  ~~ hep_job s j + hep_job s j && false <= 0

subgoal 2 (ID 2122) is:
 ~~ hep_job s j + hep_job s j && true <=
 has
   (fun j0 : Job =>
    ~~ hep_job s j0 || is_interference_from_another_hep_job j0 t)
   (task_arrivals_before arr_seq (job_task j) upp)

----------------------------------------------------------------------------- *)


        + rewrite Bool.andb_false_r leqn0 addn0 eqb0.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2136)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = true
  ============================
  ~~ ~~ hep_job s j

subgoal 2 (ID 2122) is:
 ~~ hep_job s j + hep_job s j && true <=
 has
   (fun j0 : Job =>
    ~~ hep_job s j0 || is_interference_from_another_hep_job j0 t)
   (task_arrivals_before arr_seq (job_task j) upp)

----------------------------------------------------------------------------- *)


          apply/negP; intros NEQ; move: NCOMPL ⇒ /negP NCOMPL; apply: NCOMPL.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2192)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = true
  NEQ : ~~ hep_job s j
  ============================
  job_completed_by j R

subgoal 2 (ID 2122) is:
 ~~ hep_job s j + hep_job s j && true <=
 has
   (fun j0 : Job =>
    ~~ hep_job s j0 || is_interference_from_another_hep_job j0 t)
   (task_arrivals_before arr_seq (job_task j) upp)

----------------------------------------------------------------------------- *)


          apply completion_monotonic with t; [ by apply ltnW | ].

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2198)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = true
  NEQ : ~~ hep_job s j
  ============================
  completed_by sched j t

subgoal 2 (ID 2122) is:
 ~~ hep_job s j + hep_job s j && true <=
 has
   (fun j0 : Job =>
    ~~ hep_job s j0 || is_interference_from_another_hep_job j0 t)
   (task_arrivals_before arr_seq (job_task j) upp)

----------------------------------------------------------------------------- *)


          apply/negP; intros NCOMPL; move: NCOMPL ⇒ /negP NCOMPL.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2284)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = true
  NEQ : ~~ hep_job s j
  NCOMPL : ~~ completed_by sched j t
  ============================
  False

subgoal 2 (ID 2122) is:
 ~~ hep_job s j + hep_job s j && true <=
 has
   (fun j0 : Job =>
    ~~ hep_job s j0 || is_interference_from_another_hep_job j0 t)
   (task_arrivals_before arr_seq (job_task j) upp)

----------------------------------------------------------------------------- *)


          move: NEQ ⇒ /negP NEQ; apply: NEQ; apply H_JLFP_respects_sequential_tasks; eauto 2.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2320)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = true
  NCOMPL : ~~ completed_by sched j t
  ============================
  job_arrival s <= job_arrival j

subgoal 2 (ID 2122) is:
 ~~ hep_job s j + hep_job s j && true <=
 has
   (fun j0 : Job =>
    ~~ hep_job s j0 || is_interference_from_another_hep_job j0 t)
   (task_arrivals_before arr_seq (job_task j) upp)

----------------------------------------------------------------------------- *)


          by eapply scheduler_executes_job_with_earliest_arrival; eauto 2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2122)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = false
  ============================
  ~~ hep_job s j + hep_job s j && true <=
  has
    (fun j0 : Job =>
     ~~ hep_job s j0 || is_interference_from_another_hep_job j0 t)
    (task_arrivals_before arr_seq (job_task j) upp)

----------------------------------------------------------------------------- *)


        + have NEQ: s != j.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2366)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = false
  ============================
  s != j

subgoal 2 (ID 2368) is:
 ~~ hep_job s j + hep_job s j && true <=
 has
   (fun j0 : Job =>
    ~~ hep_job s j0 || is_interference_from_another_hep_job j0 t)
   (task_arrivals_before arr_seq (job_task j) upp)

----------------------------------------------------------------------------- *)


          {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2366)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = false
  ============================
  s != j

----------------------------------------------------------------------------- *)


apply/negP; intros EQ2; move: EQ2 ⇒ /eqP EQ2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2426)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = false
  EQ2 : s = j
  ============================
  False

----------------------------------------------------------------------------- *)


              by move: TSKEQ ⇒ /eqP TSKEQ; apply: TSKEQ; rewrite EQ2.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2368)

subgoal 1 (ID 2368) is:
 ~~ hep_job s j + hep_job s j && true <=
 has
   (fun j0 : Job =>
    ~~ hep_job s j0 || is_interference_from_another_hep_job j0 t)
   (task_arrivals_before arr_seq (job_task j) upp)

----------------------------------------------------------------------------- *)


}
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2368)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = false
  NEQ : s != j
  ============================
  ~~ hep_job s j + hep_job s j && true <=
  has
    (fun j0 : Job =>
     ~~ hep_job s j0 || is_interference_from_another_hep_job j0 t)
    (task_arrivals_before arr_seq (job_task j) upp)

----------------------------------------------------------------------------- *)



          have Fact: b, ~~ b + b = true; first by intros b; destruct b.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2476)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = false
  NEQ : s != j
  Fact : forall b : bool, ~~ b + b = true
  ============================
  ~~ hep_job s j + hep_job s j && true <=
  has
    (fun j0 : Job =>
     ~~ hep_job s j0 || is_interference_from_another_hep_job j0 t)
    (task_arrivals_before arr_seq (job_task j) upp)

----------------------------------------------------------------------------- *)


          rewrite Bool.andb_true_r Fact; simpl; rewrite lt0b; clear Fact.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2494)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = false
  NEQ : s != j
  ============================
  has
    (fun j0 : Job =>
     ~~ hep_job s j0 || is_interference_from_another_hep_job j0 t)
    (task_arrivals_before arr_seq (job_task j) upp)

----------------------------------------------------------------------------- *)


          apply/hasP; j.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2524)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = false
  NEQ : s != j
  ============================
  j \in task_arrivals_before arr_seq (job_task j) upp

subgoal 2 (ID 2525) is:
 ~~ hep_job s j || is_interference_from_another_hep_job j t

----------------------------------------------------------------------------- *)


          × rewrite mem_filter; apply/andP; split; first by done.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2558)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = false
  NEQ : s != j
  ============================
  j \in arrival_sequence.arrivals_between arr_seq 0 upp

subgoal 2 (ID 2525) is:
 ~~ hep_job s j || is_interference_from_another_hep_job j t

----------------------------------------------------------------------------- *)


              by eapply arrivals_between_sub with (t2 := 0) (t3 := upp); eauto 2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2525)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = false
  NEQ : s != j
  ============================
  ~~ hep_job s j || is_interference_from_another_hep_job j t

----------------------------------------------------------------------------- *)


          × destruct (hep_job s j) eqn:HP; apply/orP; [right|left]; last by done.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2635)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  j : Job
  t1 : nat
  R, upp : instant
  ARRin : arrives_in arr_seq j
  TSK : job_task j = tsk
  ARR : j \in arrivals_before arr_seq upp
  NCOMPL : ~~ job_completed_by j R
  t : nat
  LT' : t < R
  s : Job
  Sched_s : scheduled_at sched s t
  EqSched_s : #|[pred x |
                  let
                  'FiniteQuant.Quantified F :=
                   FiniteQuant.ex (T:=Core) (, scheduled_on s (sched t) x) x
                     x in F]| <> 0
  EQ : sched t = Some s
  TSKEQ : (job_task s == job_task j) = false
  NEQ : s != j
  HP : hep_job s j = true
  ============================
  is_interference_from_another_hep_job j t

----------------------------------------------------------------------------- *)


              by rewrite /is_interference_from_another_hep_job EQ
                         /another_hep_job NEQ Bool.andb_true_r.
(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


    Qed.

In this section we prove that the (abstract) cumulative interfering workload is equivalent to conventional workload, i.e., the one defined with concrete schedule parameters.
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.
      Lemma instantiated_cumulative_workload_of_hep_jobs_equal_total_workload_of_hep_jobs:
        cumulative_interfering_workload_of_hep_jobs j t1 t2
        = workload_of_other_hep_jobs j t1 t2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1326)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  t1, t2 : instant
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  ============================
  cumulative_interfering_workload_of_hep_jobs j t1 t2 =
  workload_of_other_hep_jobs j t1 t2

----------------------------------------------------------------------------- *)


      Proof.
        intros.
        rewrite /cumulative_interfering_workload_of_hep_jobs
                /workload_of_other_hep_jobs.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1328)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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) interfering_workload_of_hep_jobs j t =
  workload_of_jobs (another_hep_job^~ j) (arrivals_between t1 t2)

----------------------------------------------------------------------------- *)


        case NEQ: (t1 < t2); last first.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1416)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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) interfering_workload_of_hep_jobs j t =
  workload_of_jobs (another_hep_job^~ j) (arrivals_between t1 t2)

subgoal 2 (ID 1376) is:
 \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
 workload_of_jobs (another_hep_job^~ j) (arrivals_between t1 t2)

----------------------------------------------------------------------------- *)


        {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1416)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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) interfering_workload_of_hep_jobs j t =
  workload_of_jobs (another_hep_job^~ j) (arrivals_between t1 t2)

----------------------------------------------------------------------------- *)


move: NEQ ⇒ /negP /negP; rewrite -leqNgt; moveNEQ.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1499)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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) interfering_workload_of_hep_jobs j t =
  workload_of_jobs (another_hep_job^~ j) (arrivals_between t1 t2)

----------------------------------------------------------------------------- *)


          rewrite big_geq; last by done.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1509)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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 t1 t2)

----------------------------------------------------------------------------- *)


          rewrite /arrivals_between /arrival_sequence.arrivals_between big_geq; last by done.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1524)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1376)

subgoal 1 (ID 1376) is:
 \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
 workload_of_jobs (another_hep_job^~ j) (arrivals_between t1 t2)

----------------------------------------------------------------------------- *)


        }

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1376)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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) interfering_workload_of_hep_jobs j t =
  workload_of_jobs (another_hep_job^~ j) (arrivals_between t1 t2)

----------------------------------------------------------------------------- *)


        {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1376)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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) interfering_workload_of_hep_jobs j t =
  workload_of_jobs (another_hep_job^~ j) (arrivals_between t1 t2)

----------------------------------------------------------------------------- *)


unfold interfering_workload_of_hep_jobs, workload_of_jobs.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1540)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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 t1 t2 | another_hep_job j0 j) job_cost j0

----------------------------------------------------------------------------- *)


          convert_two_instants_into_instant_and_duration t1 t2 k.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1567)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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 t1 (t1 + k) | another_hep_job j0 j)
     job_cost j0

----------------------------------------------------------------------------- *)


          induction k.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1571)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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 t1 (t1 + 0) | another_hep_job j0 j)
     job_cost j0

subgoal 2 (ID 1574) is:
 \sum_(t1 <= t < t1 + succn k)
    \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp =
 \sum_(j0 <- arrivals_between t1 (t1 + succn k) | 
 another_hep_job j0 j) job_cost j0

----------------------------------------------------------------------------- *)


          - rewrite !addn0.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1579)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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 t1 t1 | another_hep_job j0 j) job_cost j0

subgoal 2 (ID 1574) is:
 \sum_(t1 <= t < t1 + succn k)
    \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp =
 \sum_(j0 <- arrivals_between t1 (t1 + succn k) | 
 another_hep_job j0 j) job_cost j0

----------------------------------------------------------------------------- *)


            rewrite big_geq; last by done.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1592)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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 t1 t1 | another_hep_job j0 j) job_cost j0

subgoal 2 (ID 1574) is:
 \sum_(t1 <= t < t1 + succn k)
    \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp =
 \sum_(j0 <- arrivals_between t1 (t1 + succn k) | 
 another_hep_job j0 j) job_cost j0

----------------------------------------------------------------------------- *)


            rewrite /arrivals_between /arrival_sequence.arrivals_between big_geq; last by done.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1607)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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

subgoal 2 (ID 1574) is:
 \sum_(t1 <= t < t1 + succn k)
    \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp =
 \sum_(j0 <- arrivals_between t1 (t1 + succn k) | 
 another_hep_job j0 j) job_cost j0

----------------------------------------------------------------------------- *)


              by rewrite /workload_of_jobs big_nil.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1574)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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 t1 (t1 + k) | 
        another_hep_job j0 j) job_cost j0
  ============================
  \sum_(t1 <= t < t1 + succn k)
     \sum_(jhp <- arrivals_at arr_seq t | another_hep_job jhp j) job_cost jhp =
  \sum_(j0 <- arrivals_between t1 (t1 + succn k) | 
  another_hep_job j0 j) job_cost j0

----------------------------------------------------------------------------- *)


          - rewrite addnS big_nat_recr //=; last by rewrite leq_addr.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1636)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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 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 t1 (succn (t1 + k)) | 
  another_hep_job j0 j) job_cost j0

----------------------------------------------------------------------------- *)


            rewrite IHk.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1687)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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 t1 (t1 + k) | 
        another_hep_job j0 j) job_cost j0
  ============================
  \sum_(j0 <- arrivals_between 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 t1 (succn (t1 + k)) | 
  another_hep_job j0 j) job_cost j0

----------------------------------------------------------------------------- *)


            rewrite /arrivals_between /arrival_sequence.arrivals_between big_nat_recr //=;
                    last by rewrite leq_addr.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1703)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  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 t1 (t1 + k) | 
        another_hep_job j0 j) job_cost j0
  ============================
  \sum_(j0 <- \cat_(t1<=t<t1 + k|true)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|true)arrivals_at arr_seq i ++
               arrivals_at arr_seq (t1 + k)) | another_hep_job j0 j)
     job_cost j0

----------------------------------------------------------------------------- *)


              by rewrite big_cat //=.

(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


        }

(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


      Qed.

    End InstantiatedWorkloadEquivalence.

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].
Same for the two notions of a busy interval.
In this section we prove that the (abstract) cumulative interference of jobs with higher or equal priority is equal to total service of jobs with higher or equal priority.
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.

We consider an arbitrary time interval [t1, t) that starts with a quiet time.
      Variable t1 t : instant.
      Hypothesis H_quiet_time : quiet_time_cl j t1.

Then for any job j, the (abstract) instantiated function of interference is equal to the total service of jobs with higher or equal priority.
      Lemma instantiated_cumulative_interference_of_hep_jobs_equal_total_interference_of_hep_jobs:
        cumulative_interference_from_other_hep_jobs j t1 t
        = service_of_other_hep_jobs j t1 t.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1348)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  ============================
  cumulative_interference_from_other_hep_jobs j t1 t =
  service_of_other_hep_jobs j t1 t

----------------------------------------------------------------------------- *)


      Proof.
        rewrite /service_of_other_hep_jobs /service_of_jobs.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1356)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  ============================
  cumulative_interference_from_other_hep_jobs j t1 t =
  \sum_(j0 <- arrivals_between t1 t | another_hep_job j0 j)
     service_during sched j0 t1 t

----------------------------------------------------------------------------- *)


        rewrite /cumulative_interference_from_other_hep_jobs
                /is_interference_from_another_hep_job.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1358)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  ============================
  \sum_(t1 <= t0 < t)
     match sched t0 with
     | Some jhp => another_hep_job jhp j
     | None => false
     end =
  \sum_(j0 <- arrivals_between t1 t | another_hep_job j0 j)
     service_during sched j0 t1 t

----------------------------------------------------------------------------- *)


        rewrite exchange_big //= big_nat_cond [in X in _ = X]big_nat_cond.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1419)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  ============================
  \sum_(t1 <= i < t | (t1 <= i < t) && true)
     match sched i with
     | Some jhp => another_hep_job jhp j
     | None => false
     end =
  \sum_(t1 <= i < t | (t1 <= i < t) && true)
     \sum_(i0 <- arrivals_between t1 t | another_hep_job i0 j)
        (sched i == Some i0)

----------------------------------------------------------------------------- *)


        all: apply/eqP; rewrite eqn_leq; apply/andP; split.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1504)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  ============================
  \sum_(t1 <= i < t | (t1 <= i < t) && true)
     match sched i with
     | Some jhp => another_hep_job jhp j
     | None => false
     end <=
  \sum_(t1 <= i < t | (t1 <= i < t) && true)
     \sum_(i0 <- arrivals_between t1 t | another_hep_job i0 j)
        (sched i == Some i0)

subgoal 2 (ID 1505) is:
 \sum_(t1 <= i < t | (t1 <= i < t) && true)
    \sum_(i0 <- arrivals_between t1 t | another_hep_job i0 j)
       (sched i == Some i0) <=
 \sum_(t1 <= i < t | (t1 <= i < t) && true)
    match sched i with
    | Some jhp => another_hep_job jhp j
    | None => false
    end

----------------------------------------------------------------------------- *)


        all: rewrite leq_sum //; movex /andP [/andP [Ge Le] _].

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1654)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  ============================
  match sched x with
  | Some jhp => another_hep_job jhp j
  | None => false
  end <=
  \sum_(i <- arrivals_between t1 t | another_hep_job i j) (sched x == Some i)

subgoal 2 (ID 1733) is:
 \sum_(i <- arrivals_between t1 t | another_hep_job i j) (sched x == Some i) <=
 match sched x with
 | Some jhp => another_hep_job jhp j
 | None => false
 end

----------------------------------------------------------------------------- *)


        {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1654)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  ============================
  match sched x with
  | Some jhp => another_hep_job jhp j
  | None => false
  end <=
  \sum_(i <- arrivals_between t1 t | another_hep_job i j) (sched x == Some i)

----------------------------------------------------------------------------- *)


ideal_proc_model_sched_case_analysis_eq sched x jo; first by done.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1829)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  ============================
  match sched x with
  | Some jhp => another_hep_job jhp j
  | None => false
  end <=
  \sum_(i <- arrivals_between t1 t | another_hep_job i j) (sched x == Some i)

----------------------------------------------------------------------------- *)


          move: (Sched_jo); rewrite scheduled_at_def; move ⇒ /eqP EQ; rewrite EQ.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1874)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  ============================
  another_hep_job jo j <=
  \sum_(i <- arrivals_between t1 t | another_hep_job i j) (Some jo == Some i)

----------------------------------------------------------------------------- *)


          destruct (another_hep_job jo j) eqn:PRIO; last by done.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1886)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : another_hep_job jo j = true
  ============================
  true <=
  \sum_(i <- arrivals_between t1 t | another_hep_job i j) (Some jo == Some i)

----------------------------------------------------------------------------- *)


          rewrite (big_rem jo) //=; first by rewrite PRIO eq_refl.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1928)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : another_hep_job jo j = true
  ============================
  jo \in arrivals_between t1 t

----------------------------------------------------------------------------- *)


          apply arrived_between_implies_in_arrivals; try done.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1961)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : another_hep_job jo j = true
  ============================
  arrives_in arr_seq jo

subgoal 2 (ID 1962) is:
 arrived_between jo t1 t

----------------------------------------------------------------------------- *)


          - by apply H_jobs_come_from_arrival_sequence with x.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1962)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : another_hep_job jo j = true
  ============================
  arrived_between jo t1 t

----------------------------------------------------------------------------- *)


          - rewrite /arrived_between; apply/andP; split.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2045)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : another_hep_job jo j = true
  ============================
  t1 <= job_arrival jo

subgoal 2 (ID 2046) is:
 job_arrival jo < t

----------------------------------------------------------------------------- *)


            + move: PRIO ⇒ /andP [PRIO1 PRIO2].

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2088)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO1 : hep_job jo j
  PRIO2 : jo != j
  ============================
  t1 <= job_arrival jo

subgoal 2 (ID 2046) is:
 job_arrival jo < t

----------------------------------------------------------------------------- *)


              rewrite leqNgt; apply/negP; intros AB.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2114)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO1 : hep_job jo j
  PRIO2 : jo != j
  AB : job_arrival jo < t1
  ============================
  False

subgoal 2 (ID 2046) is:
 job_arrival jo < t

----------------------------------------------------------------------------- *)


              move: (Sched_jo).
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2115)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO1 : hep_job jo j
  PRIO2 : jo != j
  AB : job_arrival jo < t1
  ============================
  scheduled_at sched jo x -> False

subgoal 2 (ID 2046) is:
 job_arrival jo < t

----------------------------------------------------------------------------- *)


rewrite -[scheduled_at _ _ _]Bool.negb_involutive;
                                  move ⇒ /negP SCHED2; apply: SCHED2.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2159)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO1 : hep_job jo j
  PRIO2 : jo != j
  AB : job_arrival jo < t1
  ============================
  ~~ scheduled_at sched jo x

subgoal 2 (ID 2046) is:
 job_arrival jo < t

----------------------------------------------------------------------------- *)


              apply completed_implies_not_scheduled; eauto.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2165)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO1 : hep_job jo j
  PRIO2 : jo != j
  AB : job_arrival jo < t1
  ============================
  completed_by sched jo x

subgoal 2 (ID 2046) is:
 job_arrival jo < t

----------------------------------------------------------------------------- *)


              apply completion_monotonic with t1; try done.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 3102)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO1 : hep_job jo j
  PRIO2 : jo != j
  AB : job_arrival jo < t1
  ============================
  completed_by sched jo t1

subgoal 2 (ID 2046) is:
 job_arrival jo < t

----------------------------------------------------------------------------- *)


              apply H_quiet_time; try done.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 3126)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO1 : hep_job jo j
  PRIO2 : jo != j
  AB : job_arrival jo < t1
  ============================
  arrives_in arr_seq jo

subgoal 2 (ID 2046) is:
 job_arrival jo < t

----------------------------------------------------------------------------- *)


              eapply H_jobs_come_from_arrival_sequence; eauto.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2046)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : another_hep_job jo j = true
  ============================
  job_arrival jo < t

----------------------------------------------------------------------------- *)


            + by apply leq_ltn_trans with x; [apply H_jobs_must_arrive_to_execute | done].

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1733)

subgoal 1 (ID 1733) is:
 \sum_(i <- arrivals_between t1 t | another_hep_job i j) (sched x == Some i) <=
 match sched x with
 | Some jhp => another_hep_job jhp j
 | None => false
 end

----------------------------------------------------------------------------- *)


        }

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1733)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  ============================
  \sum_(i <- arrivals_between t1 t | another_hep_job i j) (sched x == Some i) <=
  match sched x with
  | Some jhp => another_hep_job jhp j
  | None => false
  end

----------------------------------------------------------------------------- *)


        {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1733)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  ============================
  \sum_(i <- arrivals_between t1 t | another_hep_job i j) (sched x == Some i) <=
  match sched x with
  | Some jhp => another_hep_job jhp j
  | None => false
  end

----------------------------------------------------------------------------- *)


ideal_proc_model_sched_case_analysis_eq sched x jo; first by rewrite big1_eq.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 3259)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  ============================
  \sum_(i <- arrivals_between t1 t | another_hep_job i j) (sched x == Some i) <=
  match sched x with
  | Some jhp => another_hep_job jhp j
  | None => false
  end

----------------------------------------------------------------------------- *)


          move: (Sched_jo); rewrite scheduled_at_def; move ⇒ /eqP EQ; rewrite EQ.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 3312)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  ============================
  \sum_(i <- arrivals_between t1 t | another_hep_job i j) (Some jo == Some i) <=
  another_hep_job jo j

----------------------------------------------------------------------------- *)


          destruct (another_hep_job jo j) eqn:PRIO.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 3324)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : another_hep_job jo j = true
  ============================
  \sum_(i <- arrivals_between t1 t | another_hep_job i j) (Some jo == Some i) <=
  true

subgoal 2 (ID 3325) is:
 \sum_(i <- arrivals_between t1 t | another_hep_job i j) (Some jo == Some i) <=
 false

----------------------------------------------------------------------------- *)


          - rewrite -EQ.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 3327)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : another_hep_job jo j = true
  ============================
  \sum_(i <- arrivals_between t1 t | another_hep_job i j) (sched x == Some i) <=
  true

subgoal 2 (ID 3325) is:
 \sum_(i <- arrivals_between t1 t | another_hep_job i j) (Some jo == Some i) <=
 false

----------------------------------------------------------------------------- *)


have SCH := service_of_jobs_le_1 sched _ (arrivals_between t1 t) _ x.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 3337)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : another_hep_job jo j = true
  SCH : forall b : Job -> bool,
        uniq (arrivals_between t1 t) ->
        \sum_(j <- arrivals_between t1 t | b j) service_at sched j x <= 1
  ============================
  \sum_(i <- arrivals_between t1 t | another_hep_job i j) (sched x == Some i) <=
  true

subgoal 2 (ID 3325) is:
 \sum_(i <- arrivals_between t1 t | another_hep_job i j) (Some jo == Some i) <=
 false

----------------------------------------------------------------------------- *)


            eapply leq_trans; last by apply SCH; eauto using arrivals_uniq.

(* ----------------------------------[ coqtop ]---------------------------------

2 focused subgoals
(shelved: 1) (ID 3339)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : another_hep_job jo j = true
  SCH : forall b : Job -> bool,
        uniq (arrivals_between t1 t) ->
        \sum_(j <- arrivals_between t1 t | b j) service_at sched j x <= 1
  ============================
  \sum_(i <- arrivals_between t1 t | another_hep_job i j) (sched x == Some i) <=
  \sum_(j0 <- arrivals_between t1 t | ?b j0) service_at sched j0 x

subgoal 2 (ID 3325) is:
 \sum_(i <- arrivals_between t1 t | another_hep_job i j) (Some jo == Some i) <=
 false

----------------------------------------------------------------------------- *)


            erewrite leq_sum; eauto 2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 3325)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : another_hep_job jo j = false
  ============================
  \sum_(i <- arrivals_between t1 t | another_hep_job i j) (Some jo == Some i) <=
  false

----------------------------------------------------------------------------- *)


          - rewrite leqn0 big1 //; intros joo PRIO2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 4035)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : another_hep_job jo j = false
  joo : Job
  PRIO2 : another_hep_job joo j
  ============================
  (Some jo == Some joo) = 0

----------------------------------------------------------------------------- *)


            apply/eqP; rewrite eqb0; apply/eqP; intros C.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 4133)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : another_hep_job jo j = false
  joo : Job
  PRIO2 : another_hep_job joo j
  C : Some jo = Some joo
  ============================
  False

----------------------------------------------------------------------------- *)


            inversion C; subst joo; clear C.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 4159)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : another_hep_job jo j = false
  PRIO2 : another_hep_job jo j
  ============================
  False

----------------------------------------------------------------------------- *)


              by rewrite PRIO2 in PRIO.

(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


        }

(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


      Qed.

The same applies to the alternative definition of interference.
      Lemma instantiated_cumulative_interference_of_hep_tasks_equal_total_interference_of_hep_tasks:
        cumulative_interference_from_hep_jobs_from_other_tasks j t1 t
        = service_of_hep_jobs_from_other_tasks j t1 t.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1351)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  ============================
  cumulative_interference_from_hep_jobs_from_other_tasks j t1 t =
  service_of_hep_jobs_from_other_tasks j t1 t

----------------------------------------------------------------------------- *)


      Proof.
        rewrite /service_of_hep_jobs_from_other_tasks /service_of_jobs.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1359)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  ============================
  cumulative_interference_from_hep_jobs_from_other_tasks j t1 t =
  \sum_(j0 <- arrivals_between t1 t | hep_job_from_another_task j0 j)
     service_during sched j0 t1 t

----------------------------------------------------------------------------- *)


        rewrite /cumulative_interference_from_hep_jobs_from_other_tasks
                /is_interference_from_hep_job_from_another_task.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1361)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  ============================
  \sum_(t1 <= t0 < t)
     match sched t0 with
     | Some jhp => hep_job_from_another_task jhp j
     | None => false
     end =
  \sum_(j0 <- arrivals_between t1 t | hep_job_from_another_task j0 j)
     service_during sched j0 t1 t

----------------------------------------------------------------------------- *)


        rewrite exchange_big //= big_nat_cond [in X in _ = X]big_nat_cond.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1422)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  ============================
  \sum_(t1 <= i < t | (t1 <= i < t) && true)
     match sched i with
     | Some jhp => hep_job_from_another_task jhp j
     | None => false
     end =
  \sum_(t1 <= i < t | (t1 <= i < t) && true)
     \sum_(i0 <- arrivals_between t1 t | hep_job_from_another_task i0 j)
        (sched i == Some i0)

----------------------------------------------------------------------------- *)


        all: apply/eqP; rewrite eqn_leq; apply/andP; split.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1507)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  ============================
  \sum_(t1 <= i < t | (t1 <= i < t) && true)
     match sched i with
     | Some jhp => hep_job_from_another_task jhp j
     | None => false
     end <=
  \sum_(t1 <= i < t | (t1 <= i < t) && true)
     \sum_(i0 <- arrivals_between t1 t | hep_job_from_another_task i0 j)
        (sched i == Some i0)

subgoal 2 (ID 1508) is:
 \sum_(t1 <= i < t | (t1 <= i < t) && true)
    \sum_(i0 <- arrivals_between t1 t | hep_job_from_another_task i0 j)
       (sched i == Some i0) <=
 \sum_(t1 <= i < t | (t1 <= i < t) && true)
    match sched i with
    | Some jhp => hep_job_from_another_task jhp j
    | None => false
    end

----------------------------------------------------------------------------- *)


        all: rewrite leq_sum //; movex /andP [/andP [Ge Le] _].

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1657)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  ============================
  match sched x with
  | Some jhp => hep_job_from_another_task jhp j
  | None => false
  end <=
  \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
     (sched x == Some i)

subgoal 2 (ID 1736) is:
 \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
    (sched x == Some i) <=
 match sched x with
 | Some jhp => hep_job_from_another_task jhp j
 | None => false
 end

----------------------------------------------------------------------------- *)


        {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1657)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  ============================
  match sched x with
  | Some jhp => hep_job_from_another_task jhp j
  | None => false
  end <=
  \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
     (sched x == Some i)

----------------------------------------------------------------------------- *)


ideal_proc_model_sched_case_analysis_eq sched x jo; first by done.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1832)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  ============================
  match sched x with
  | Some jhp => hep_job_from_another_task jhp j
  | None => false
  end <=
  \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
     (sched x == Some i)

----------------------------------------------------------------------------- *)


          move: (Sched_jo); rewrite scheduled_at_def; move ⇒ /eqP EQ; rewrite EQ.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1877)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  ============================
  hep_job_from_another_task jo j <=
  \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
     (Some jo == Some i)

----------------------------------------------------------------------------- *)


          destruct (hep_job_from_another_task jo j) eqn:PRIO; last by done.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1889)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : hep_job_from_another_task jo j = true
  ============================
  true <=
  \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
     (Some jo == Some i)

----------------------------------------------------------------------------- *)


          rewrite (big_rem jo) //=; first by rewrite PRIO eq_refl.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1931)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : hep_job_from_another_task jo j = true
  ============================
  jo \in arrivals_between t1 t

----------------------------------------------------------------------------- *)


          apply arrived_between_implies_in_arrivals; try done.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1964)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : hep_job_from_another_task jo j = true
  ============================
  arrives_in arr_seq jo

subgoal 2 (ID 1965) is:
 arrived_between jo t1 t

----------------------------------------------------------------------------- *)


          - by apply H_jobs_come_from_arrival_sequence with x.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1965)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : hep_job_from_another_task jo j = true
  ============================
  arrived_between jo t1 t

----------------------------------------------------------------------------- *)


          - rewrite /arrived_between; apply/andP; split.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2048)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : hep_job_from_another_task jo j = true
  ============================
  t1 <= job_arrival jo

subgoal 2 (ID 2049) is:
 job_arrival jo < t

----------------------------------------------------------------------------- *)


            + move: PRIO ⇒ /andP [PRIO1 PRIO2].

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2091)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO1 : hep_job jo j
  PRIO2 : job_task jo != job_task j
  ============================
  t1 <= job_arrival jo

subgoal 2 (ID 2049) is:
 job_arrival jo < t

----------------------------------------------------------------------------- *)


              rewrite leqNgt; apply/negP; intros AB.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2117)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO1 : hep_job jo j
  PRIO2 : job_task jo != job_task j
  AB : job_arrival jo < t1
  ============================
  False

subgoal 2 (ID 2049) is:
 job_arrival jo < t

----------------------------------------------------------------------------- *)


              move: (Sched_jo).
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2118)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO1 : hep_job jo j
  PRIO2 : job_task jo != job_task j
  AB : job_arrival jo < t1
  ============================
  scheduled_at sched jo x -> False

subgoal 2 (ID 2049) is:
 job_arrival jo < t

----------------------------------------------------------------------------- *)


rewrite -[scheduled_at _ _ _]Bool.negb_involutive;
                                  move ⇒ /negP SCHED2; apply: SCHED2.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2162)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO1 : hep_job jo j
  PRIO2 : job_task jo != job_task j
  AB : job_arrival jo < t1
  ============================
  ~~ scheduled_at sched jo x

subgoal 2 (ID 2049) is:
 job_arrival jo < t

----------------------------------------------------------------------------- *)


              apply completed_implies_not_scheduled; eauto.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2168)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO1 : hep_job jo j
  PRIO2 : job_task jo != job_task j
  AB : job_arrival jo < t1
  ============================
  completed_by sched jo x

subgoal 2 (ID 2049) is:
 job_arrival jo < t

----------------------------------------------------------------------------- *)


              apply completion_monotonic with t1; try done.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 3105)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO1 : hep_job jo j
  PRIO2 : job_task jo != job_task j
  AB : job_arrival jo < t1
  ============================
  completed_by sched jo t1

subgoal 2 (ID 2049) is:
 job_arrival jo < t

----------------------------------------------------------------------------- *)


              apply H_quiet_time; try done.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 3129)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO1 : hep_job jo j
  PRIO2 : job_task jo != job_task j
  AB : job_arrival jo < t1
  ============================
  arrives_in arr_seq jo

subgoal 2 (ID 2049) is:
 job_arrival jo < t

----------------------------------------------------------------------------- *)


              eapply H_jobs_come_from_arrival_sequence; simpl; eauto.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2049)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : hep_job_from_another_task jo j = true
  ============================
  job_arrival jo < t

----------------------------------------------------------------------------- *)


            + by apply leq_ltn_trans with x; [apply H_jobs_must_arrive_to_execute | done].

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1736)

subgoal 1 (ID 1736) is:
 \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
    (sched x == Some i) <=
 match sched x with
 | Some jhp => hep_job_from_another_task jhp j
 | None => false
 end

----------------------------------------------------------------------------- *)


        }

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1736)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  ============================
  \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
     (sched x == Some i) <=
  match sched x with
  | Some jhp => hep_job_from_another_task jhp j
  | None => false
  end

----------------------------------------------------------------------------- *)


        {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1736)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  ============================
  \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
     (sched x == Some i) <=
  match sched x with
  | Some jhp => hep_job_from_another_task jhp j
  | None => false
  end

----------------------------------------------------------------------------- *)


ideal_proc_model_sched_case_analysis_eq sched x jo; first by rewrite big1_eq.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 3264)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  ============================
  \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
     (sched x == Some i) <=
  match sched x with
  | Some jhp => hep_job_from_another_task jhp j
  | None => false
  end

----------------------------------------------------------------------------- *)


          move: (Sched_jo); rewrite scheduled_at_def; move ⇒ /eqP EQ; rewrite EQ.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 3317)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  ============================
  \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
     (Some jo == Some i) <= hep_job_from_another_task jo j

----------------------------------------------------------------------------- *)


          destruct (hep_job_from_another_task jo j) eqn:PRIO.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 3329)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : hep_job_from_another_task jo j = true
  ============================
  \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
     (Some jo == Some i) <= true

subgoal 2 (ID 3330) is:
 \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
    (Some jo == Some i) <= false

----------------------------------------------------------------------------- *)


          - rewrite -EQ.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 3332)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : hep_job_from_another_task jo j = true
  ============================
  \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
     (sched x == Some i) <= true

subgoal 2 (ID 3330) is:
 \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
    (Some jo == Some i) <= false

----------------------------------------------------------------------------- *)


have SCH := service_of_jobs_le_1 sched _ (arrivals_between t1 t) _ x.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 3342)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : hep_job_from_another_task jo j = true
  SCH : forall b : Job -> bool,
        uniq (arrivals_between t1 t) ->
        \sum_(j <- arrivals_between t1 t | b j) service_at sched j x <= 1
  ============================
  \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
     (sched x == Some i) <= true

subgoal 2 (ID 3330) is:
 \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
    (Some jo == Some i) <= false

----------------------------------------------------------------------------- *)


            eapply leq_trans; last by apply SCH; eauto using arrivals_uniq.

(* ----------------------------------[ coqtop ]---------------------------------

2 focused subgoals
(shelved: 1) (ID 3344)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : hep_job_from_another_task jo j = true
  SCH : forall b : Job -> bool,
        uniq (arrivals_between t1 t) ->
        \sum_(j <- arrivals_between t1 t | b j) service_at sched j x <= 1
  ============================
  \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
     (sched x == Some i) <=
  \sum_(j0 <- arrivals_between t1 t | ?b j0) service_at sched j0 x

subgoal 2 (ID 3330) is:
 \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
    (Some jo == Some i) <= false

----------------------------------------------------------------------------- *)


            erewrite leq_sum; eauto 2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 3330)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : hep_job_from_another_task jo j = false
  ============================
  \sum_(i <- arrivals_between t1 t | hep_job_from_another_task i j)
     (Some jo == Some i) <= false

----------------------------------------------------------------------------- *)


          - rewrite leqn0 big1 //; intros joo PRIO2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 4040)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : hep_job_from_another_task jo j = false
  joo : Job
  PRIO2 : hep_job_from_another_task joo j
  ============================
  (Some jo == Some joo) = 0

----------------------------------------------------------------------------- *)


            apply/eqP; rewrite eqb0; apply/eqP; intros C.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 4138)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : hep_job_from_another_task jo j = false
  joo : Job
  PRIO2 : hep_job_from_another_task joo j
  C : Some jo = Some joo
  ============================
  False

----------------------------------------------------------------------------- *)


            inversion C; subst joo; clear C.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 4164)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_of_task tsk j
  t1, t : instant
  H_quiet_time : quiet_time_cl j t1
  x : nat
  Ge : t1 <= x
  Le : x < t
  jo : Job
  Sched_jo : scheduled_at sched jo x
  EqSched_jo : #|[pred x0 |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched x) x0)
                      x0 x0 in F]| <> 0
  EQ : sched x = Some jo
  PRIO : hep_job_from_another_task jo j = false
  PRIO2 : hep_job_from_another_task jo j
  ============================
  False

----------------------------------------------------------------------------- *)


              by rewrite PRIO2 in PRIO.

(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


        }

(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


      Qed.

      End InstantiatedServiceEquivalences.

In this section we prove that the abstract definition of busy interval is equivalent to the conventional, concrete definition of busy interval for JLFP scheduling.
      Section BusyIntervalEquivalence.

Consider any job j of [tsk].
        Variable j : Job.
        Hypothesis H_j_arrives : arrives_in arr_seq j.
        Hypothesis H_job_of_tsk : job_task j = tsk.
        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.

        Lemma quiet_time_cl_implies_quiet_time_ab:
           t, quiet_time_cl j t quiet_time_ab j t.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1351)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  ============================
  forall t : instant, quiet_time_cl j t -> quiet_time_ab j t

----------------------------------------------------------------------------- *)


        Proof.
          have zero_is_quiet_time: j, quiet_time_cl j 0.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1353)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  ============================
  forall j0 : Job, quiet_time_cl j0 0

subgoal 2 (ID 1355) is:
 forall t : instant, quiet_time_cl j t -> quiet_time_ab j t

----------------------------------------------------------------------------- *)


          {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1353)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  ============================
  forall j0 : Job, quiet_time_cl j0 0

----------------------------------------------------------------------------- *)


by intros jhp ARR HP AB; move: AB; rewrite /arrived_before ltn0.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1355)

subgoal 1 (ID 1355) is:
 forall t : instant, quiet_time_cl j t -> quiet_time_ab j t

----------------------------------------------------------------------------- *)


}

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1355)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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

----------------------------------------------------------------------------- *)


          intros t QT; split.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1400)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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
  ============================
  cumul_interference interference j 0 t =
  cumul_interfering_workload interfering_workload j 0 t

subgoal 2 (ID 1401) is:
 ~~ pending_earlier_and_at sched j t

----------------------------------------------------------------------------- *)


          {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1400)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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
  ============================
  cumul_interference interference j 0 t =
  cumul_interfering_workload interfering_workload j 0 t

----------------------------------------------------------------------------- *)


intros.
            have CIS := cumulative_interference_split.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1406)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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
  CIS : forall (j : Job) (t1 t2 : instant),
        cumulative_interference j t1 t2 =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  ============================
  cumul_interference interference j 0 t =
  cumul_interfering_workload interfering_workload j 0 t

----------------------------------------------------------------------------- *)


            rewrite /cumulative_interference in CIS.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1459)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  ============================
  cumul_interference interference j 0 t =
  cumul_interfering_workload interfering_workload j 0 t

----------------------------------------------------------------------------- *)


            rewrite /cumul_interference /cumul_interfering_workload.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1465)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  ============================
  \sum_(0 <= t0 < t) interference j t0 =
  \sum_(0 <= t0 < t) interfering_workload j t0

----------------------------------------------------------------------------- *)


            rewrite CIS !big_split //=.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1489)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  ============================
  cumulative_priority_inversion j 0 t +
  cumulative_interference_from_other_hep_jobs j 0 t =
  \sum_(0 <= i < t) is_priority_inversion j i +
  \sum_(0 <= i < t) interfering_workload_of_hep_jobs j i

----------------------------------------------------------------------------- *)


            apply/eqP; rewrite eqn_add2l.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1571)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  ============================
  cumulative_interference_from_other_hep_jobs j 0 t ==
  \sum_(0 <= i < t) interfering_workload_of_hep_jobs j i

----------------------------------------------------------------------------- *)


            rewrite instantiated_cumulative_interference_of_hep_jobs_equal_total_interference_of_hep_jobs;
              last by apply zero_is_quiet_time.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1577)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  ============================
  service_of_other_hep_jobs j 0 t ==
  \sum_(0 <= i < t) interfering_workload_of_hep_jobs j i

----------------------------------------------------------------------------- *)


            have L2 := instantiated_cumulative_workload_of_hep_jobs_equal_total_workload_of_hep_jobs.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1583)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  L2 : forall (t1 t2 : instant) (j : Job),
       cumulative_interfering_workload_of_hep_jobs j t1 t2 =
       workload_of_other_hep_jobs j t1 t2
  ============================
  service_of_other_hep_jobs j 0 t ==
  \sum_(0 <= i < t) interfering_workload_of_hep_jobs j i

----------------------------------------------------------------------------- *)


            rewrite /cumulative_interfering_workload_of_hep_jobs in L2; rewrite L2; clear L2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1643)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  ============================
  service_of_other_hep_jobs j 0 t == workload_of_other_hep_jobs j 0 t

----------------------------------------------------------------------------- *)


            rewrite eq_sym; apply/eqP.
(* ----------------------------------[ coqtop ]---------------------------------

1 focused subgoal
(shelved: 1) (ID 1676)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  ============================
  workload_of_other_hep_jobs j 0 t = service_of_other_hep_jobs j 0 t

----------------------------------------------------------------------------- *)


            apply all_jobs_have_completed_equiv_workload_eq_service; try done.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1711)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  ============================
  forall j0 : Job,
  j0 \in arrival_sequence.arrivals_between arr_seq 0 t ->
  another_hep_job j0 j -> completed_by sched j0 t

----------------------------------------------------------------------------- *)


            intros; apply QT.

(* ----------------------------------[ coqtop ]---------------------------------

3 subgoals (ID 1743)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  j0 : Job
  H4 : j0 \in arrival_sequence.arrivals_between arr_seq 0 t
  H5 : another_hep_job j0 j
  ============================
  arrives_in arr_seq j0

subgoal 2 (ID 1744) is:
 hep_job j0 j
subgoal 3 (ID 1745) is:
 arrived_before j0 t

----------------------------------------------------------------------------- *)


            - by apply in_arrivals_implies_arrived in H4.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1744)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  j0 : Job
  H4 : j0 \in arrival_sequence.arrivals_between arr_seq 0 t
  H5 : another_hep_job j0 j
  ============================
  hep_job j0 j

subgoal 2 (ID 1745) is:
 arrived_before j0 t

----------------------------------------------------------------------------- *)


            - by move: H5 ⇒ /andP [H6 H7].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1745)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  j0 : Job
  H4 : j0 \in arrival_sequence.arrivals_between arr_seq 0 t
  H5 : another_hep_job j0 j
  ============================
  arrived_before j0 t

----------------------------------------------------------------------------- *)


            - by apply in_arrivals_implies_arrived_between in H4.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1401)

subgoal 1 (ID 1401) is:
 ~~ pending_earlier_and_at sched j t

----------------------------------------------------------------------------- *)


          }

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1401)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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

----------------------------------------------------------------------------- *)


          {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1401)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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

----------------------------------------------------------------------------- *)


rewrite negb_and Bool.negb_involutive; apply/orP.

(* ----------------------------------[ coqtop ]---------------------------------

1 focused subgoal
(shelved: 1) (ID 1826)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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

----------------------------------------------------------------------------- *)


            case ARR: (arrived_before j t); [right | by left].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1930)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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
  ARR : arrived_before j t = true
  ============================
  completed_by sched j t

----------------------------------------------------------------------------- *)


            apply QT; try eauto 2.

(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


          }

(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


        Qed.

        Lemma quiet_time_ab_implies_quiet_time_cl:
           t, quiet_time_ab j t quiet_time_cl j t.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1353)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  ============================
  forall t : instant, quiet_time_ab j t -> quiet_time_cl j t

----------------------------------------------------------------------------- *)


        Proof.
          have zero_is_quiet_time: j, quiet_time_cl j 0.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1355)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  ============================
  forall j0 : Job, quiet_time_cl j0 0

subgoal 2 (ID 1357) is:
 forall t : instant, quiet_time_ab j t -> quiet_time_cl j t

----------------------------------------------------------------------------- *)


          {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1355)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  ============================
  forall j0 : Job, quiet_time_cl j0 0

----------------------------------------------------------------------------- *)


by intros jhp ARR HP AB; move: AB; rewrite /arrived_before ltn0.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1357)

subgoal 1 (ID 1357) is:
 forall t : instant, quiet_time_ab j t -> quiet_time_cl j t

----------------------------------------------------------------------------- *)


}

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1357)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  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

----------------------------------------------------------------------------- *)


          have CIS := cumulative_interference_split.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1403)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        cumulative_interference j t1 t2 =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  ============================
  forall t : instant, quiet_time_ab j t -> quiet_time_cl j t

----------------------------------------------------------------------------- *)


          have IC1 := instantiated_cumulative_interference_of_hep_jobs_equal_total_interference_of_hep_jobs.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1408)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        cumulative_interference j t1 t2 =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_other_hep_jobs j t1 t
  ============================
  forall t : instant, quiet_time_ab j t -> quiet_time_cl j t

----------------------------------------------------------------------------- *)


          rewrite /cumulative_interference /service_of_other_hep_jobs in CIS, IC1.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1462)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  ============================
  forall t : instant, quiet_time_ab j t -> quiet_time_cl j t

----------------------------------------------------------------------------- *)


          intros t [T0 T1]; intros jhp ARR HP ARB.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1473)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  T0 : cumul_interference interference j 0 t =
       cumul_interfering_workload interfering_workload j 0 t
  T1 : ~~ pending_earlier_and_at sched j t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  ============================
  completed_by sched jhp t

----------------------------------------------------------------------------- *)


          eapply all_jobs_have_completed_equiv_workload_eq_service with
              (P := fun jhphep_job jhp j) (t1 := 0) (t2 := t);
            eauto 2; last eapply arrived_between_implies_in_arrivals; try done.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1504)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  T0 : cumul_interference interference j 0 t =
       cumul_interfering_workload 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)
    (arrival_sequence.arrivals_between arr_seq 0 t) =
  service_of_jobs sched (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) 0 t

----------------------------------------------------------------------------- *)


          move: T0; rewrite /cumul_interference /cumul_interfering_workload.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1554)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  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
  ============================
  \sum_(0 <= t0 < t) interference j t0 =
  \sum_(0 <= t0 < t) interfering_workload j t0 ->
  workload_of_jobs (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) =
  service_of_jobs sched (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) 0 t

----------------------------------------------------------------------------- *)


          rewrite CIS !big_split //=; move ⇒ /eqP; rewrite eqn_add2l.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1665)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  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_interference_from_other_hep_jobs j 0 t ==
  \sum_(0 <= i < t) interfering_workload_of_hep_jobs j i ->
  workload_of_jobs (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) =
  service_of_jobs sched (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) 0 t

----------------------------------------------------------------------------- *)


          rewrite IC1; last by apply zero_is_quiet_time.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1671)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  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
  ============================
  service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
  \sum_(0 <= i < t) interfering_workload_of_hep_jobs j i ->
  workload_of_jobs (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) =
  service_of_jobs sched (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) 0 t

----------------------------------------------------------------------------- *)


          have L2 := instantiated_cumulative_workload_of_hep_jobs_equal_total_workload_of_hep_jobs;
                       rewrite /cumulative_interfering_workload_of_hep_jobs in L2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1736)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  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
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  ============================
  service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
  \sum_(0 <= i < t) interfering_workload_of_hep_jobs j i ->
  workload_of_jobs (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) =
  service_of_jobs sched (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) 0 t

----------------------------------------------------------------------------- *)


          rewrite L2.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1741)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  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
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  ============================
  service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
  workload_of_other_hep_jobs j 0 t ->
  workload_of_jobs (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) =
  service_of_jobs sched (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) 0 t

----------------------------------------------------------------------------- *)


moveT2.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1742)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  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
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  ============================
  workload_of_jobs (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) =
  service_of_jobs sched (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) 0 t

----------------------------------------------------------------------------- *)


apply/eqP; rewrite eq_sym.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1802)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  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
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  ============================
  service_of_jobs sched (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) 0 t ==
  workload_of_jobs (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t)

----------------------------------------------------------------------------- *)


          move: T1; rewrite negb_and Bool.negb_involutive -leqNgt; move ⇒ /orP [T1 | T1].

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1858)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  T1 : t <= job_arrival j
  ============================
  service_of_jobs sched (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) 0 t ==
  workload_of_jobs (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t)

subgoal 2 (ID 1859) is:
 service_of_jobs sched (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t) 0 t ==
 workload_of_jobs (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t)

----------------------------------------------------------------------------- *)


          - have NOTIN: j \notin arrivals_between 0 t.

(* ----------------------------------[ coqtop ]---------------------------------

3 subgoals (ID 1864)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  T1 : t <= job_arrival j
  ============================
  j \notin arrivals_between 0 t

subgoal 2 (ID 1866) is:
 service_of_jobs sched (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t) 0 t ==
 workload_of_jobs (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t)
subgoal 3 (ID 1859) is:
 service_of_jobs sched (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t) 0 t ==
 workload_of_jobs (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t)

----------------------------------------------------------------------------- *)


            {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1864)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  T1 : t <= job_arrival j
  ============================
  j \notin arrivals_between 0 t

----------------------------------------------------------------------------- *)


apply/memPn; intros jo IN; apply/negP; intros EQ; move: EQ ⇒ /eqP EQ.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1955)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  T1 : t <= job_arrival j
  jo : Job
  IN : jo \in arrivals_between 0 t
  EQ : jo = j
  ============================
  False

----------------------------------------------------------------------------- *)


              subst jo.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1961)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  T1 : t <= job_arrival j
  IN : j \in arrivals_between 0 t
  ============================
  False

----------------------------------------------------------------------------- *)


              apply in_arrivals_implies_arrived_between in IN; try done.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1964)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  T1 : t <= job_arrival j
  IN : arrived_between j 0 t
  ============================
  False

----------------------------------------------------------------------------- *)


                by move: IN ⇒ /andP [_ IN]; move: T1; rewrite leqNgt; move ⇒ /negP LT; apply: LT.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1866)

subgoal 1 (ID 1866) is:
 service_of_jobs sched (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t) 0 t ==
 workload_of_jobs (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t)
subgoal 2 (ID 1859) is:
 service_of_jobs sched (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t) 0 t ==
 workload_of_jobs (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t)

----------------------------------------------------------------------------- *)


}

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1866)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  T1 : t <= job_arrival j
  NOTIN : j \notin arrivals_between 0 t
  ============================
  service_of_jobs sched (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) 0 t ==
  workload_of_jobs (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t)

subgoal 2 (ID 1859) is:
 service_of_jobs sched (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t) 0 t ==
 workload_of_jobs (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t)

----------------------------------------------------------------------------- *)


            rewrite /workload_of_other_hep_jobs in T2.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2133)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T1 : t <= job_arrival j
  NOTIN : j \notin arrivals_between 0 t
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_jobs (another_hep_job^~ j) (arrivals_between 0 t)
  ============================
  service_of_jobs sched (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) 0 t ==
  workload_of_jobs (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t)

subgoal 2 (ID 1859) is:
 service_of_jobs sched (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t) 0 t ==
 workload_of_jobs (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t)

----------------------------------------------------------------------------- *)


              by rewrite /service_of_jobs /workload_of_jobs !sum_notin_rem_eqn in T2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1859)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  T1 : completed_by sched j t
  ============================
  service_of_jobs sched (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) 0 t ==
  workload_of_jobs (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t)

----------------------------------------------------------------------------- *)


          - have JIN: j \in arrivals_between 0 t.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2356)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  T1 : completed_by sched j t
  ============================
  j \in arrivals_between 0 t

subgoal 2 (ID 2358) is:
 service_of_jobs sched (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t) 0 t ==
 workload_of_jobs (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t)

----------------------------------------------------------------------------- *)


            {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2356)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  T1 : completed_by sched j t
  ============================
  j \in arrivals_between 0 t

----------------------------------------------------------------------------- *)


eapply completed_implies_scheduled_before in T1; eauto 2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2364)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  T1 : exists t' : nat, job_arrival j <= t' < t /\ scheduled_at sched j t'
  ============================
  j \in arrivals_between 0 t

----------------------------------------------------------------------------- *)


              apply arrived_between_implies_in_arrivals; try done.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2384)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  T1 : exists t' : nat, job_arrival j <= t' < t /\ scheduled_at sched j t'
  ============================
  arrived_between j 0 t

----------------------------------------------------------------------------- *)


              move: T1 ⇒ [t' [T3 _]].

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2430)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  t' : nat
  T3 : job_arrival j <= t' < t
  ============================
  arrived_between j 0 t

----------------------------------------------------------------------------- *)


              apply/andP; split; first by done.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2457)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  t' : nat
  T3 : job_arrival j <= t' < t
  ============================
  job_arrival j < t

----------------------------------------------------------------------------- *)


              move: T3 ⇒ /andP [H3e H3t].

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2499)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  t' : nat
  H3e : job_arrival j <= t'
  H3t : t' < t
  ============================
  job_arrival j < t

----------------------------------------------------------------------------- *)


                by apply leq_ltn_trans with t'.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2358)

subgoal 1 (ID 2358) is:
 service_of_jobs sched (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t) 0 t ==
 workload_of_jobs (hep_job^~ j)
   (arrival_sequence.arrivals_between arr_seq 0 t)

----------------------------------------------------------------------------- *)


}

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2358)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  T1 : completed_by sched j t
  JIN : j \in arrivals_between 0 t
  ============================
  service_of_jobs sched (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) 0 t ==
  workload_of_jobs (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t)

----------------------------------------------------------------------------- *)


            have UNIC: uniq (arrivals_between 0 t) by eapply arrivals_uniq; eauto 2.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2509)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  T1 : completed_by sched j t
  JIN : j \in arrivals_between 0 t
  UNIC : uniq (arrivals_between 0 t)
  ============================
  service_of_jobs sched (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t) 0 t ==
  workload_of_jobs (hep_job^~ j)
    (arrival_sequence.arrivals_between arr_seq 0 t)

----------------------------------------------------------------------------- *)


            unfold service_of_jobs, workload_of_jobs.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2511)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  T1 : completed_by sched j t
  JIN : j \in arrivals_between 0 t
  UNIC : uniq (arrivals_between 0 t)
  ============================
  \sum_(j0 <- arrival_sequence.arrivals_between arr_seq 0 t | 
  hep_job j0 j) service_during sched j0 0 t ==
  \sum_(j0 <- arrival_sequence.arrivals_between arr_seq 0 t | 
  hep_job j0 j) job_cost j0

----------------------------------------------------------------------------- *)


            rewrite big_mkcond //= (bigD1_seq j) //= -big_mkcondl //=.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2594)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T2 : service_of_jobs sched (another_hep_job^~ j) (arrivals_between 0 t) 0 t ==
       workload_of_other_hep_jobs j 0 t
  T1 : completed_by sched j t
  JIN : j \in arrivals_between 0 t
  UNIC : uniq (arrivals_between 0 t)
  ============================
  (if hep_job j j then service_during sched j 0 t else 0) +
  \sum_(i <- arrival_sequence.arrivals_between arr_seq 0 t | 
  hep_job i j && (i != j)) service_during sched i 0 t ==
  \sum_(j0 <- arrival_sequence.arrivals_between arr_seq 0 t | 
  hep_job j0 j) job_cost j0

----------------------------------------------------------------------------- *)


            move: T2; rewrite /service_of_jobs; move ⇒ /eqP T2; rewrite T2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2663)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T1 : completed_by sched j t
  JIN : j \in arrivals_between 0 t
  UNIC : uniq (arrivals_between 0 t)
  T2 : \sum_(j0 <- arrivals_between 0 t | another_hep_job j0 j)
          service_during sched j0 0 t = workload_of_other_hep_jobs j 0 t
  ============================
  (if hep_job j j then service_during sched j 0 t else 0) +
  workload_of_other_hep_jobs j 0 t ==
  \sum_(j0 <- arrival_sequence.arrivals_between arr_seq 0 t | 
  hep_job j0 j) job_cost j0

----------------------------------------------------------------------------- *)


            rewrite [X in _ == X]big_mkcond //= [X in _ == X](bigD1_seq j) //= -big_mkcondl //=.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2762)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T1 : completed_by sched j t
  JIN : j \in arrivals_between 0 t
  UNIC : uniq (arrivals_between 0 t)
  T2 : \sum_(j0 <- arrivals_between 0 t | another_hep_job j0 j)
          service_during sched j0 0 t = workload_of_other_hep_jobs j 0 t
  ============================
  (if hep_job j j then service_during sched j 0 t else 0) +
  workload_of_other_hep_jobs j 0 t ==
  (if hep_job j j then job_cost j else 0) +
  \sum_(i <- arrival_sequence.arrivals_between arr_seq 0 t | 
  hep_job i j && (i != j)) job_cost i

----------------------------------------------------------------------------- *)


            rewrite eqn_add2r; unfold hep_job.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2793)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T1 : completed_by sched j t
  JIN : j \in arrivals_between 0 t
  UNIC : uniq (arrivals_between 0 t)
  T2 : \sum_(j0 <- arrivals_between 0 t | another_hep_job j0 j)
          service_during sched j0 0 t = workload_of_other_hep_jobs j 0 t
  ============================
  (if H3 j j then service_during sched j 0 t else 0) ==
  (if H3 j j then job_cost j else 0)

----------------------------------------------------------------------------- *)


            erewrite H_priority_is_reflexive; eauto 2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2843)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T1 : completed_by sched j t
  JIN : j \in arrivals_between 0 t
  UNIC : uniq (arrivals_between 0 t)
  T2 : \sum_(j0 <- arrivals_between 0 t | another_hep_job j0 j)
          service_during sched j0 0 t = workload_of_other_hep_jobs j 0 t
  ============================
  service_during sched j 0 t == job_cost j

----------------------------------------------------------------------------- *)


            rewrite eqn_leq; apply/andP; split; try eauto 2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2882)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  zero_is_quiet_time : forall j : Job, quiet_time_cl j 0
  CIS : forall (j : Job) (t1 t2 : instant),
        \sum_(t1 <= t < t2) interference j t =
        cumulative_priority_inversion j t1 t2 +
        cumulative_interference_from_other_hep_jobs j t1 t2
  IC1 : forall (j : Job) (t1 t : instant),
        quiet_time_cl j t1 ->
        cumulative_interference_from_other_hep_jobs j t1 t =
        service_of_jobs sched (another_hep_job^~ j) 
          (arrivals_between t1 t) t1 t
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARB : arrived_before jhp t
  L2 : forall (t1 t2 : instant) (j : Job),
       \sum_(t1 <= t < t2) interfering_workload_of_hep_jobs j t =
       workload_of_other_hep_jobs j t1 t2
  T1 : completed_by sched j t
  JIN : j \in arrivals_between 0 t
  UNIC : uniq (arrivals_between 0 t)
  T2 : \sum_(j0 <- arrivals_between 0 t | another_hep_job j0 j)
          service_during sched j0 0 t = workload_of_other_hep_jobs j 0 t
  ============================
  service_during sched j 0 t <= job_cost j

----------------------------------------------------------------------------- *)


              by apply service_at_most_cost; eauto with basic_facts.

(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


        Qed.

The equivalence trivially follows from the lemmas above.
        Corollary instantiated_quiet_time_equivalent_quiet_time:
           t,
            quiet_time_cl j t quiet_time_ab j t.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1355)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  ============================
  forall t : instant, quiet_time_cl j t <-> quiet_time_ab j t

----------------------------------------------------------------------------- *)


        Proof.
          intros ?; split.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1358)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t : instant
  ============================
  quiet_time_cl j t -> quiet_time_ab j t

subgoal 2 (ID 1359) is:
 quiet_time_ab j t -> quiet_time_cl j t

----------------------------------------------------------------------------- *)


          - apply quiet_time_cl_implies_quiet_time_ab.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1359)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t : instant
  ============================
  quiet_time_ab j t -> quiet_time_cl j t

----------------------------------------------------------------------------- *)


          - apply quiet_time_ab_implies_quiet_time_cl.

(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


        Qed.

Based on that, we prove that the concept of busy interval obtained by instantiating the abstract definition of busy interval coincides with the conventional definition of busy interval.
        Lemma instantiated_busy_interval_equivalent_edf_busy_interval:
           t1 t2,
            busy_interval_cl j t1 t2 busy_interval_ab j t1 t2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1358)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  ============================
  forall t1 t2 : instant,
  busy_interval_cl j t1 t2 <-> busy_interval_ab j t1 t2

----------------------------------------------------------------------------- *)


        Proof.
          split.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1362)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  ============================
  busy_interval_cl j t1 t2 -> busy_interval_ab j t1 t2

subgoal 2 (ID 1363) is:
 busy_interval_ab j t1 t2 -> busy_interval_cl j t1 t2

----------------------------------------------------------------------------- *)


          {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1362)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  ============================
  busy_interval_cl j t1 t2 -> busy_interval_ab j t1 t2

----------------------------------------------------------------------------- *)


move ⇒ [[NEQ [QTt1 [NQT REL]] QTt2]].

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1404)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  NEQ : t1 < t2
  QTt1 : busy_interval.quiet_time arr_seq sched j t1
  NQT : forall t : nat,
        t1 < t < t2 -> ~ busy_interval.quiet_time arr_seq sched j t
  REL : t1 <= job_arrival j < t2
  QTt2 : busy_interval.quiet_time arr_seq sched j t2
  ============================
  busy_interval_ab j t1 t2

----------------------------------------------------------------------------- *)


            split; [split; [ |split] | ].

(* ----------------------------------[ coqtop ]---------------------------------

4 subgoals (ID 1409)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  NEQ : t1 < t2
  QTt1 : busy_interval.quiet_time arr_seq sched j t1
  NQT : forall t : nat,
        t1 < t < t2 -> ~ busy_interval.quiet_time arr_seq sched j t
  REL : t1 <= job_arrival j < t2
  QTt2 : busy_interval.quiet_time arr_seq sched j t2
  ============================
  t1 <= job_arrival j < t2

subgoal 2 (ID 1412) is:
 quiet_time sched interference interfering_workload j t1
subgoal 3 (ID 1413) is:
 forall t : nat,
 t1 < t < t2 -> ~ quiet_time sched interference interfering_workload j t
subgoal 4 (ID 1407) is:
 quiet_time sched interference interfering_workload j t2

----------------------------------------------------------------------------- *)


            - by done.

(* ----------------------------------[ coqtop ]---------------------------------

3 subgoals (ID 1412)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  NEQ : t1 < t2
  QTt1 : busy_interval.quiet_time arr_seq sched j t1
  NQT : forall t : nat,
        t1 < t < t2 -> ~ busy_interval.quiet_time arr_seq sched j t
  REL : t1 <= job_arrival j < t2
  QTt2 : busy_interval.quiet_time arr_seq sched j t2
  ============================
  quiet_time sched interference interfering_workload j t1

subgoal 2 (ID 1413) is:
 forall t : nat,
 t1 < t < t2 -> ~ quiet_time sched interference interfering_workload j t
subgoal 3 (ID 1407) is:
 quiet_time sched interference interfering_workload j t2

----------------------------------------------------------------------------- *)


            - by apply instantiated_quiet_time_equivalent_quiet_time in QTt1.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1413)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  NEQ : t1 < t2
  QTt1 : busy_interval.quiet_time arr_seq sched j t1
  NQT : forall t : nat,
        t1 < t < t2 -> ~ busy_interval.quiet_time arr_seq sched j t
  REL : t1 <= job_arrival j < t2
  QTt2 : busy_interval.quiet_time arr_seq sched j t2
  ============================
  forall t : nat,
  t1 < t < t2 -> ~ quiet_time sched interference interfering_workload j t

subgoal 2 (ID 1407) is:
 quiet_time sched interference interfering_workload j t2

----------------------------------------------------------------------------- *)


            - by intros t NE QT; eapply NQT; eauto 2; apply instantiated_quiet_time_equivalent_quiet_time.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1407)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  NEQ : t1 < t2
  QTt1 : busy_interval.quiet_time arr_seq sched j t1
  NQT : forall t : nat,
        t1 < t < t2 -> ~ busy_interval.quiet_time arr_seq sched j t
  REL : t1 <= job_arrival j < t2
  QTt2 : busy_interval.quiet_time arr_seq sched j t2
  ============================
  quiet_time sched interference interfering_workload j t2

----------------------------------------------------------------------------- *)


            - by eapply instantiated_quiet_time_equivalent_quiet_time in QTt2.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1363)

subgoal 1 (ID 1363) is:
 busy_interval_ab j t1 t2 -> busy_interval_cl j t1 t2

----------------------------------------------------------------------------- *)


          }

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1363)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  ============================
  busy_interval_ab j t1 t2 -> busy_interval_cl j t1 t2

----------------------------------------------------------------------------- *)


          {
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1363)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  ============================
  busy_interval_ab j t1 t2 -> busy_interval_cl j t1 t2

----------------------------------------------------------------------------- *)


move ⇒ [[/andP [NEQ1 NEQ2] [QTt1 NQT] QTt2]].

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1510)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  NEQ1 : t1 <= job_arrival j
  NEQ2 : job_arrival j < t2
  QTt1 : quiet_time sched interference interfering_workload j t1
  NQT : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QTt2 : quiet_time sched interference interfering_workload j t2
  ============================
  busy_interval_cl j t1 t2

----------------------------------------------------------------------------- *)


            split; [split; [ | split; [ |split] ] | ].

(* ----------------------------------[ coqtop ]---------------------------------

5 subgoals (ID 1515)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  NEQ1 : t1 <= job_arrival j
  NEQ2 : job_arrival j < t2
  QTt1 : quiet_time sched interference interfering_workload j t1
  NQT : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QTt2 : quiet_time sched interference interfering_workload j t2
  ============================
  t1 < t2

subgoal 2 (ID 1518) is:
 busy_interval.quiet_time arr_seq sched j t1
subgoal 3 (ID 1521) is:
 forall t : nat, t1 < t < t2 -> ~ busy_interval.quiet_time arr_seq sched j t
subgoal 4 (ID 1522) is:
 t1 <= job_arrival j < t2
subgoal 5 (ID 1513) is:
 busy_interval.quiet_time arr_seq sched j t2

----------------------------------------------------------------------------- *)


            - by apply leq_ltn_trans with (job_arrival j).

(* ----------------------------------[ coqtop ]---------------------------------

4 subgoals (ID 1518)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  NEQ1 : t1 <= job_arrival j
  NEQ2 : job_arrival j < t2
  QTt1 : quiet_time sched interference interfering_workload j t1
  NQT : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QTt2 : quiet_time sched interference interfering_workload j t2
  ============================
  busy_interval.quiet_time arr_seq sched j t1

subgoal 2 (ID 1521) is:
 forall t : nat, t1 < t < t2 -> ~ busy_interval.quiet_time arr_seq sched j t
subgoal 3 (ID 1522) is:
 t1 <= job_arrival j < t2
subgoal 4 (ID 1513) is:
 busy_interval.quiet_time arr_seq sched j t2

----------------------------------------------------------------------------- *)


            - by eapply instantiated_quiet_time_equivalent_quiet_time; eauto 2.

(* ----------------------------------[ coqtop ]---------------------------------

3 subgoals (ID 1521)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  NEQ1 : t1 <= job_arrival j
  NEQ2 : job_arrival j < t2
  QTt1 : quiet_time sched interference interfering_workload j t1
  NQT : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QTt2 : quiet_time sched interference interfering_workload j t2
  ============================
  forall t : nat, t1 < t < t2 -> ~ busy_interval.quiet_time arr_seq sched j t

subgoal 2 (ID 1522) is:
 t1 <= job_arrival j < t2
subgoal 3 (ID 1513) is:
 busy_interval.quiet_time arr_seq sched j t2

----------------------------------------------------------------------------- *)


            - by intros t NEQ QT; eapply NQT; eauto 2; eapply instantiated_quiet_time_equivalent_quiet_time in QT; eauto 2.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1522)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  NEQ1 : t1 <= job_arrival j
  NEQ2 : job_arrival j < t2
  QTt1 : quiet_time sched interference interfering_workload j t1
  NQT : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QTt2 : quiet_time sched interference interfering_workload j t2
  ============================
  t1 <= job_arrival j < t2

subgoal 2 (ID 1513) is:
 busy_interval.quiet_time arr_seq sched j t2

----------------------------------------------------------------------------- *)


            - by apply/andP; split.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1513)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H_sequential_tasks : sequential_tasks arr_seq sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H_JLFP_respects_sequential_tasks : policy_respects_sequential_tasks
  tsk : Task
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  cumulative_task_interference := cumul_task_interference arr_seq sched
   : (Job -> instant -> bool) -> Task -> instant -> nat -> nat -> nat
  another_hep_job := fun j1 j2 : Job => hep_job j1 j2 && (j1 != j2)
   : JLFP_policy Job
  hep_job_from_another_task := fun j1 j2 : Job =>
                               hep_job j1 j2 && (job_task j1 != job_task j2)
   : JLFP_policy Job
  is_priority_inversion := fun j : Job =>
                           [eta priority_inversion.is_priority_inversion
                                  sched j] : Job -> instant -> bool
  cumulative_interference_from_other_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 is_interference_from_another_hep_job
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interference := fun (j : Job) (t1 t2 : instant) =>
                             \sum_(t1 <= t < t2) interference j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload_of_hep_jobs := fun 
                                                 (j : Job) 
                                                 (t1 t2 : instant) =>
                                                 \sum_(t1 <= t < t2)
                                                  
                                                 interfering_workload_of_hep_jobs
                                                 j t
   : Job -> instant -> instant -> nat
  cumulative_interfering_workload := fun (j : Job) (t1 t2 : instant) =>
                                     \sum_(t1 <= t < t2)
                                        interfering_workload j t
   : Job -> instant -> instant -> nat
  workload_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                                workload_of_jobs (another_hep_job^~ j)
                                  (arrivals_between t1 t2)
   : Job -> instant -> instant -> nat
  service_of_hep_jobs_from_other_tasks := fun (j : Job) (t1 t2 : instant) =>
                                          service_of_jobs sched
                                            (hep_job_from_another_task^~ j)
                                            (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  service_of_other_hep_jobs := fun (j : Job) (t1 t2 : instant) =>
                               service_of_jobs sched 
                                 (another_hep_job^~ j)
                                 (arrivals_between t1 t2) t1 t2
   : Job -> instant -> instant -> nat
  quiet_time_cl := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  quiet_time_ab := quiet_time sched interference interfering_workload
   : Job -> instant -> Prop
  busy_interval_cl := busy_interval.busy_interval arr_seq sched
   : Job -> instant -> instant -> Prop
  busy_interval_ab := busy_interval sched interference interfering_workload
   : Job -> instant -> instant -> Prop
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  NEQ1 : t1 <= job_arrival j
  NEQ2 : job_arrival j < t2
  QTt1 : quiet_time sched interference interfering_workload j t1
  NQT : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QTt2 : quiet_time sched interference interfering_workload j t2
  ============================
  busy_interval.quiet_time arr_seq sched j t2

----------------------------------------------------------------------------- *)


            - by eapply instantiated_quiet_time_equivalent_quiet_time; eauto 2.

(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


          }

(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


        Qed.

      End BusyIntervalEquivalence.

  End Equivalences.

End JLFPInstantiation.