Library prosa.analysis.facts.busy_interval.busy_interval


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

Welcome to Coq 8.11.2 (June 2020)

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


Require Export prosa.model.schedule.work_conserving.
Require Export prosa.analysis.definitions.job_properties.
Require Export prosa.analysis.definitions.priority_inversion.
Require Export prosa.analysis.facts.behavior.all.
Require Export prosa.analysis.facts.model.service_of_jobs.

Throughout this file, we assume ideal uni-processor schedules.
Require Import prosa.model.processor.ideal.

Throughout this file, we assume the basic (i.e., Liu & Layland) readiness model.
Require Import prosa.model.readiness.basic.

Existence of Busy Interval for JLFP-models

In this module we derive a sufficient condition for existence of busy intervals for uni-processor for JLFP schedulers.
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 a given JLFP policy.
  Context `{JLFP_policy Job}.

For simplicity, let's define some local names.
Consider an arbitrary task [tsk].
  Variable tsk : Task.

Consider an arbitrary job [j].
  Variable j : Job.
  Hypothesis H_from_arrival_sequence : arrives_in arr_seq j.
  Hypothesis H_job_task : job_of_task tsk j.
  Hypothesis H_job_cost_positive : job_cost_positive j.

Recall the list of jobs that arrive in any interval.
We begin by proving a basic lemma about completion of the job within its busy interval.
  Section BasicLemma.

Assume that the priority relation is reflexive.
Consider any busy interval [t1, t2) of job [j].
    Variable t1 t2 : instant.
    Hypothesis H_busy_interval : busy_interval t1 t2.

We prove that job j completes by the end of the busy interval.
    Lemma job_completes_within_busy_interval:
      job_completed_by j t2.

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

1 subgoal (ID 1882)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_priority_is_reflexive : reflexive_priorities
  t1, t2 : instant
  H_busy_interval : busy_interval t1 t2
  ============================
  job_completed_by j t2

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


    Proof.
      rename H_priority_is_reflexive into REFL, H_busy_interval into BUSY.

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

1 subgoal (ID 1883)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  REFL : reflexive_priorities
  t1, t2 : instant
  BUSY : busy_interval t1 t2
  ============================
  job_completed_by j t2

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


      move: BUSY ⇒ [[_ [_ [_ /andP [_ ARR]]]] QUIET].

(* ----------------------------------[ 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  REFL : reflexive_priorities
  t1, t2 : instant
  ARR : job_arrival j < t2
  QUIET : busy_interval.quiet_time arr_seq sched j t2
  ============================
  job_completed_by j t2

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


      apply QUIET; try done.

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

1 subgoal (ID 1967)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  REFL : reflexive_priorities
  t1, t2 : instant
  ARR : job_arrival j < t2
  QUIET : busy_interval.quiet_time arr_seq sched j t2
  ============================
  hep_job j j

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


      apply (REFL 0).

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

No more subgoals.

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


    Qed.

  End BasicLemma.

In this section, we prove that during a busy interval there always exists a pending job.
  Section ExistsPendingJob.

Let [[t1, t2]] be any interval where time t1 is quiet and time t2 is not quiet.
    Variable t1 t2 : instant.
    Hypothesis H_interval : t1 t2.
    Hypothesis H_quiet : quiet_time t1.
    Hypothesis H_not_quiet : ¬ quiet_time t2.

Then, we prove that there is a job pending at time t2 that has higher or equal priority (with respect of [tsk]).
    Lemma not_quiet_implies_exists_pending_job:
       j_hp,
        arrives_in arr_seq j_hp
        arrived_between j_hp t1 t2
        hep_job j_hp j
        ¬ job_completed_by j_hp t2.
(* ----------------------------------[ 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t1, t2 : instant
  H_interval : t1 <= t2
  H_quiet : quiet_time t1
  H_not_quiet : ~ quiet_time t2
  ============================
  exists j_hp : Job,
    arrives_in arr_seq j_hp /\
    arrived_between j_hp t1 t2 /\
    hep_job j_hp j /\ ~ job_completed_by j_hp t2

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


    Proof.
      rename H_quiet into QUIET, H_not_quiet into NOTQUIET.

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

1 subgoal (ID 1887)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t1, t2 : instant
  H_interval : t1 <= t2
  QUIET : quiet_time t1
  NOTQUIET : ~ quiet_time t2
  ============================
  exists j_hp : Job,
    arrives_in arr_seq j_hp /\
    arrived_between j_hp t1 t2 /\
    hep_job j_hp j /\ ~ job_completed_by j_hp t2

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


      destruct (has (fun j_hp(~~ job_completed_by j_hp t2) && hep_job j_hp j)
                    (arrivals_between t1 t2)) eqn:COMP.

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

2 subgoals (ID 1901)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t1, t2 : instant
  H_interval : t1 <= t2
  QUIET : quiet_time t1
  NOTQUIET : ~ quiet_time t2
  COMP : has
           (fun j_hp : Job => ~~ job_completed_by j_hp t2 && hep_job j_hp j)
           (arrivals_between t1 t2) = true
  ============================
  exists j_hp : Job,
    arrives_in arr_seq j_hp /\
    arrived_between j_hp t1 t2 /\
    hep_job j_hp j /\ ~ job_completed_by j_hp t2

subgoal 2 (ID 1902) is:
 exists j_hp : Job,
   arrives_in arr_seq j_hp /\
   arrived_between j_hp t1 t2 /\ hep_job j_hp j /\ ~ job_completed_by j_hp t2

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


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

1 subgoal (ID 1901)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t1, t2 : instant
  H_interval : t1 <= t2
  QUIET : quiet_time t1
  NOTQUIET : ~ quiet_time t2
  COMP : has
           (fun j_hp : Job => ~~ job_completed_by j_hp t2 && hep_job j_hp j)
           (arrivals_between t1 t2) = true
  ============================
  exists j_hp : Job,
    arrives_in arr_seq j_hp /\
    arrived_between j_hp t1 t2 /\
    hep_job j_hp j /\ ~ job_completed_by j_hp t2

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


move: COMP ⇒ /hasP [j_hp ARR /andP [NOTCOMP HP]].

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

1 subgoal (ID 1988)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t1, t2 : instant
  H_interval : t1 <= t2
  QUIET : quiet_time t1
  NOTQUIET : ~ quiet_time t2
  j_hp : Job
  ARR : j_hp \in arrivals_between t1 t2
  NOTCOMP : ~~ job_completed_by j_hp t2
  HP : hep_job j_hp j
  ============================
  exists j_hp0 : Job,
    arrives_in arr_seq j_hp0 /\
    arrived_between j_hp0 t1 t2 /\
    hep_job j_hp0 j /\ ~ job_completed_by j_hp0 t2

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


        move: (ARR) ⇒ INarr.

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

1 subgoal (ID 1990)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t1, t2 : instant
  H_interval : t1 <= t2
  QUIET : quiet_time t1
  NOTQUIET : ~ quiet_time t2
  j_hp : Job
  ARR : j_hp \in arrivals_between t1 t2
  NOTCOMP : ~~ job_completed_by j_hp t2
  HP : hep_job j_hp j
  INarr : j_hp \in arrivals_between t1 t2
  ============================
  exists j_hp0 : Job,
    arrives_in arr_seq j_hp0 /\
    arrived_between j_hp0 t1 t2 /\
    hep_job j_hp0 j /\ ~ job_completed_by j_hp0 t2

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


        apply in_arrivals_implies_arrived_between in ARR; last by done.

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

1 subgoal (ID 1993)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t1, t2 : instant
  H_interval : t1 <= t2
  QUIET : quiet_time t1
  NOTQUIET : ~ quiet_time t2
  j_hp : Job
  ARR : arrived_between j_hp t1 t2
  NOTCOMP : ~~ job_completed_by j_hp t2
  HP : hep_job j_hp j
  INarr : j_hp \in arrivals_between t1 t2
  ============================
  exists j_hp0 : Job,
    arrives_in arr_seq j_hp0 /\
    arrived_between j_hp0 t1 t2 /\
    hep_job j_hp0 j /\ ~ job_completed_by j_hp0 t2

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


        apply in_arrivals_implies_arrived in INarr.

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

1 subgoal (ID 1997)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t1, t2 : instant
  H_interval : t1 <= t2
  QUIET : quiet_time t1
  NOTQUIET : ~ quiet_time t2
  j_hp : Job
  ARR : arrived_between j_hp t1 t2
  NOTCOMP : ~~ job_completed_by j_hp t2
  HP : hep_job j_hp j
  INarr : arrives_in arr_seq j_hp
  ============================
  exists j_hp0 : Job,
    arrives_in arr_seq j_hp0 /\
    arrived_between j_hp0 t1 t2 /\
    hep_job j_hp0 j /\ ~ job_completed_by j_hp0 t2

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


          by j_hp; repeat split; last by apply/negP.

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

1 subgoal (ID 1902)

subgoal 1 (ID 1902) is:
 exists j_hp : Job,
   arrives_in arr_seq j_hp /\
   arrived_between j_hp t1 t2 /\ hep_job j_hp j /\ ~ job_completed_by j_hp t2

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


      }

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

1 subgoal (ID 1902)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t1, t2 : instant
  H_interval : t1 <= t2
  QUIET : quiet_time t1
  NOTQUIET : ~ quiet_time t2
  COMP : has
           (fun j_hp : Job => ~~ job_completed_by j_hp t2 && hep_job j_hp j)
           (arrivals_between t1 t2) = false
  ============================
  exists j_hp : Job,
    arrives_in arr_seq j_hp /\
    arrived_between j_hp t1 t2 /\
    hep_job j_hp j /\ ~ job_completed_by j_hp t2

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


      {

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

1 subgoal (ID 1902)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t1, t2 : instant
  H_interval : t1 <= t2
  QUIET : quiet_time t1
  NOTQUIET : ~ quiet_time t2
  COMP : has
           (fun j_hp : Job => ~~ job_completed_by j_hp t2 && hep_job j_hp j)
           (arrivals_between t1 t2) = false
  ============================
  exists j_hp : Job,
    arrives_in arr_seq j_hp /\
    arrived_between j_hp t1 t2 /\
    hep_job j_hp j /\ ~ job_completed_by j_hp t2

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


        apply negbT in COMP; rewrite -all_predC in COMP.

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

1 subgoal (ID 2100)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t1, t2 : instant
  H_interval : t1 <= t2
  QUIET : quiet_time t1
  NOTQUIET : ~ quiet_time t2
  COMP : all
           (predC
              (fun j_hp : Job =>
               ~~ job_completed_by j_hp t2 && hep_job j_hp j))
           (arrivals_between t1 t2)
  ============================
  exists j_hp : Job,
    arrives_in arr_seq j_hp /\
    arrived_between j_hp t1 t2 /\
    hep_job j_hp j /\ ~ job_completed_by j_hp t2

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


        move: COMP ⇒ /allP COMP.

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

1 subgoal (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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t1, t2 : instant
  H_interval : t1 <= t2
  QUIET : quiet_time t1
  NOTQUIET : ~ quiet_time t2
  COMP : {in arrivals_between t1 t2,
           forall x : Job,
           predC
             (fun j_hp : Job => ~~ job_completed_by j_hp t2 && hep_job j_hp j)
             x}
  ============================
  exists j_hp : Job,
    arrives_in arr_seq j_hp /\
    arrived_between j_hp t1 t2 /\
    hep_job j_hp j /\ ~ job_completed_by j_hp t2

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


        exfalso; apply NOTQUIET; intros j_hp IN HP ARR.

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

1 subgoal (ID 2143)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t1, t2 : instant
  H_interval : t1 <= t2
  QUIET : quiet_time t1
  NOTQUIET : ~ quiet_time t2
  COMP : {in arrivals_between t1 t2,
           forall x : Job,
           predC
             (fun j_hp : Job => ~~ job_completed_by j_hp t2 && hep_job j_hp j)
             x}
  j_hp : Job
  IN : arrives_in arr_seq j_hp
  HP : hep_job j_hp j
  ARR : arrived_before j_hp t2
  ============================
  completed_by sched j_hp t2

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


        destruct (ltnP (job_arrival j_hp) t1) as [BEFORE | AFTER];
          first by specialize (QUIET j_hp IN HP BEFORE); apply completion_monotonic with (t := t1).

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

1 subgoal (ID 2169)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t1, t2 : instant
  H_interval : t1 <= t2
  QUIET : quiet_time t1
  NOTQUIET : ~ quiet_time t2
  COMP : {in arrivals_between t1 t2,
           forall x : Job,
           predC
             (fun j_hp : Job => ~~ job_completed_by j_hp t2 && hep_job j_hp j)
             x}
  j_hp : Job
  IN : arrives_in arr_seq j_hp
  HP : hep_job j_hp j
  ARR : arrived_before j_hp t2
  AFTER : t1 <= job_arrival j_hp
  ============================
  completed_by sched j_hp t2

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


        feed (COMP j_hp).

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

2 subgoals (ID 2178)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t1, t2 : instant
  H_interval : t1 <= t2
  QUIET : quiet_time t1
  NOTQUIET : ~ quiet_time t2
  COMP : {in arrivals_between t1 t2,
           forall x : Job,
           predC
             (fun j_hp : Job => ~~ job_completed_by j_hp t2 && hep_job j_hp j)
             x}
  j_hp : Job
  IN : arrives_in arr_seq j_hp
  HP : hep_job j_hp j
  ARR : arrived_before j_hp t2
  AFTER : t1 <= job_arrival j_hp
  ============================
  j_hp \in arrivals_between t1 t2

subgoal 2 (ID 2183) is:
 completed_by sched j_hp t2

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


          by eapply arrived_between_implies_in_arrivals; eauto 1; apply/andP; split.

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

1 subgoal (ID 2183)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t1, t2 : instant
  H_interval : t1 <= t2
  QUIET : quiet_time t1
  NOTQUIET : ~ quiet_time t2
  j_hp : Job
  COMP : (fun x : Job =>
          is_true
            (predC
               (fun j_hp : Job =>
                ~~ job_completed_by j_hp t2 && hep_job j_hp j) x)) j_hp
  IN : arrives_in arr_seq j_hp
  HP : hep_job j_hp j
  ARR : arrived_before j_hp t2
  AFTER : t1 <= job_arrival j_hp
  ============================
  completed_by sched j_hp t2

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


            by rewrite /= HP andbT negbK in COMP.

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

No more subgoals.

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


      }

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

No more subgoals.

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


    Qed.

  End ExistsPendingJob.

In this section, we prove that during a busy interval the processor is never idle.
  Section ProcessorAlwaysBusy.

Assume that the schedule is work-conserving ...
... and the priority relation is reflexive and transitive.
Consider any busy interval prefix [t1, t2).
    Variable t1 t2 : instant.
    Hypothesis H_busy_interval_prefix : busy_interval_prefix t1 t2.

We prove that if the processor is idle at a time instant t, then the next time instant [t+1] will be a quiet time.
    Lemma idle_time_implies_quiet_time_at_the_next_time_instant:
       (t : instant),
        is_idle sched t
        quiet_time t.+1.

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

1 subgoal (ID 1892)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  ============================
  forall t : instant, is_idle sched t -> quiet_time (succn t)

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


    Proof.
      intros t IDLE jhp ARR HP AB.

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

1 subgoal (ID 1899)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : instant
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  AB : arrived_before jhp (succn t)
  ============================
  completed_by sched jhp (succn t)

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


      apply negbNE; apply/negP; intros NCOMP.

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

1 subgoal (ID 1922)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : instant
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  AB : arrived_before jhp (succn t)
  NCOMP : ~~ completed_by sched jhp (succn t)
  ============================
  False

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


      rewrite /arrived_before ltnS in AB.

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

1 subgoal (ID 1977)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : instant
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  NCOMP : ~~ completed_by sched jhp (succn t)
  AB : job_arrival jhp <= t
  ============================
  False

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


      move:(H_work_conserving _ t ARR) ⇒ WC.

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

1 subgoal (ID 1981)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : instant
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  NCOMP : ~~ completed_by sched jhp (succn t)
  AB : job_arrival jhp <= t
  WC : backlogged sched jhp t ->
       exists j_other : Job, scheduled_at sched j_other t
  ============================
  False

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


      feed WC.

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

2 subgoals (ID 1982)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : instant
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  NCOMP : ~~ completed_by sched jhp (succn t)
  AB : job_arrival jhp <= t
  WC : backlogged sched jhp t ->
       exists j_other : Job, scheduled_at sched j_other t
  ============================
  backlogged sched jhp t

subgoal 2 (ID 1987) is:
 False

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


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

1 subgoal (ID 1982)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : instant
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  NCOMP : ~~ completed_by sched jhp (succn t)
  AB : job_arrival jhp <= t
  WC : backlogged sched jhp t ->
       exists j_other : Job, scheduled_at sched j_other t
  ============================
  backlogged sched jhp t

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


apply/andP.
(* ----------------------------------[ coqtop ]---------------------------------

1 focused subgoal
(shelved: 1) (ID 2011)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : instant
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  NCOMP : ~~ completed_by sched jhp (succn t)
  AB : job_arrival jhp <= t
  WC : backlogged sched jhp t ->
       exists j_other : Job, scheduled_at sched j_other t
  ============================
  job_ready sched jhp t /\ ~~ scheduled_at sched jhp t

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


split.

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

2 subgoals (ID 2013)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : instant
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  NCOMP : ~~ completed_by sched jhp (succn t)
  AB : job_arrival jhp <= t
  WC : backlogged sched jhp t ->
       exists j_other : Job, scheduled_at sched j_other t
  ============================
  job_ready sched jhp t

subgoal 2 (ID 2014) is:
 ~~ scheduled_at sched jhp t

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


        - apply/negPn/negP; rewrite negb_and; intros COMP.

(* ----------------------------------[ 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : instant
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  NCOMP : ~~ completed_by sched jhp (succn t)
  AB : job_arrival jhp <= t
  WC : backlogged sched jhp t ->
       exists j_other : Job, scheduled_at sched j_other t
  COMP : ~~ has_arrived jhp t || ~~ ~~ completed_by sched jhp t
  ============================
  False

subgoal 2 (ID 2014) is:
 ~~ scheduled_at sched jhp t

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


          move: COMP ⇒ /orP; rewrite Bool.negb_involutive; move ⇒ [/negP CON|COM]; auto.

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

2 subgoals (ID 2190)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : instant
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  NCOMP : ~~ completed_by sched jhp (succn t)
  AB : job_arrival jhp <= t
  WC : backlogged sched jhp t ->
       exists j_other : Job, scheduled_at sched j_other t
  COM : completed_by sched jhp t
  ============================
  False

subgoal 2 (ID 2014) is:
 ~~ scheduled_at sched jhp t

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


          move: NCOMP ⇒ /negP NCOMP; apply: NCOMP.

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

2 subgoals (ID 2234)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : instant
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  AB : job_arrival jhp <= t
  WC : backlogged sched jhp t ->
       exists j_other : Job, scheduled_at sched j_other t
  COM : completed_by sched jhp t
  ============================
  completed_by sched jhp (succn t)

subgoal 2 (ID 2014) is:
 ~~ scheduled_at sched jhp t

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


            by apply completion_monotonic with t.

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

1 subgoal (ID 2014)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : instant
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  NCOMP : ~~ completed_by sched jhp (succn t)
  AB : job_arrival jhp <= t
  WC : backlogged sched jhp t ->
       exists j_other : Job, scheduled_at sched j_other t
  ============================
  ~~ scheduled_at sched jhp t

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


        - by move: IDLE ⇒ /eqP IDLE; rewrite /scheduled_at scheduled_in_def IDLE.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1987)

subgoal 1 (ID 1987) is:
 False

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


      }

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

1 subgoal (ID 1987)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : instant
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  NCOMP : ~~ completed_by sched jhp (succn t)
  AB : job_arrival jhp <= t
  WC : exists j_other : Job, scheduled_at sched j_other t
  ============================
  False

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


      move: IDLE WC ⇒ /eqP IDLE [jo SCHED].

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

1 subgoal (ID 2340)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : instant
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  NCOMP : ~~ completed_by sched jhp (succn t)
  AB : job_arrival jhp <= t
  IDLE : sched t = None
  jo : Job
  SCHED : scheduled_at sched jo t
  ============================
  False

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


        by rewrite scheduled_at_def IDLE in SCHED.

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

No more subgoals.

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


    Qed.

Next, we prove that at any time instant t within the busy interval there exists a job [jhp] such that (1) job [jhp] is pending at time [t] and (2) job [jhp] has higher-or-equal priority than task [tsk].
    Lemma pending_hp_job_exists:
       t,
        t1 t < t2
         jhp,
          arrives_in arr_seq jhp
          job_pending_at jhp t
          hep_job jhp j.

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

1 subgoal (ID 1899)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  ============================
  forall t : nat,
  t1 <= t < t2 ->
  exists jhp : Job,
    arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

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


    Proof.
      movet /andP [GE LT]; move: (H_busy_interval_prefix) ⇒ [_ [QTt [NQT REL]]].

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

1 subgoal (ID 1973)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  ============================
  exists jhp : Job,
    arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

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


      move: (ltngtP t1.+1 t2) ⇒ [GT|CONTR|EQ]; first last.

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

3 subgoals (ID 1998)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  EQ : succn t1 = t2
  ============================
  exists jhp : Job,
    arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

subgoal 2 (ID 1997) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j
subgoal 3 (ID 1996) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

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


      - subst t2; rewrite ltnS in LT.

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

3 subgoals (ID 2058)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 (succn t1)
  t : nat
  GE : t1 <= t
  QTt : busy_interval.quiet_time arr_seq sched j t1
  REL : t1 <= job_arrival j < succn t1
  NQT : forall t : nat,
        t1 < t < succn t1 -> ~ busy_interval.quiet_time arr_seq sched j t
  LT : t <= t1
  ============================
  exists jhp : Job,
    arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

subgoal 2 (ID 1997) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j
subgoal 3 (ID 1996) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

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


        have EQ: t1 = t by apply/eqP; rewrite eqn_leq; apply/andP; split.

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

3 subgoals (ID 2149)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 (succn t1)
  t : nat
  GE : t1 <= t
  QTt : busy_interval.quiet_time arr_seq sched j t1
  REL : t1 <= job_arrival j < succn t1
  NQT : forall t : nat,
        t1 < t < succn t1 -> ~ busy_interval.quiet_time arr_seq sched j t
  LT : t <= t1
  EQ : t1 = t
  ============================
  exists jhp : Job,
    arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

subgoal 2 (ID 1997) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j
subgoal 3 (ID 1996) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

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


        subst t1; clear GE LT.

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

3 subgoals (ID 2166)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t : nat
  NQT : forall t0 : nat,
        t < t0 < succn t -> ~ busy_interval.quiet_time arr_seq sched j t0
  REL : t <= job_arrival j < succn t
  QTt : busy_interval.quiet_time arr_seq sched j t
  H_busy_interval_prefix : busy_interval_prefix t (succn t)
  ============================
  exists jhp : Job,
    arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

subgoal 2 (ID 1997) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j
subgoal 3 (ID 1996) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

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


         j; repeat split; try done.

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

4 subgoals (ID 2174)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t : nat
  NQT : forall t0 : nat,
        t < t0 < succn t -> ~ busy_interval.quiet_time arr_seq sched j t0
  REL : t <= job_arrival j < succn t
  QTt : busy_interval.quiet_time arr_seq sched j t
  H_busy_interval_prefix : busy_interval_prefix t (succn t)
  ============================
  job_pending_at j t

subgoal 2 (ID 2175) is:
 hep_job j j
subgoal 3 (ID 1997) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j
subgoal 4 (ID 1996) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

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


        + move: REL; rewrite ltnS -eqn_leq eq_sym; move ⇒ /eqP REL.

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

4 subgoals (ID 2270)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t : nat
  NQT : forall t0 : nat,
        t < t0 < succn t -> ~ busy_interval.quiet_time arr_seq sched j t0
  QTt : busy_interval.quiet_time arr_seq sched j t
  H_busy_interval_prefix : busy_interval_prefix t (succn t)
  REL : job_arrival j = t
  ============================
  job_pending_at j t

subgoal 2 (ID 2175) is:
 hep_job j j
subgoal 3 (ID 1997) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j
subgoal 4 (ID 1996) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

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


            by rewrite -REL; eapply job_pending_at_arrival; eauto 2.

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

3 subgoals (ID 2175)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t : nat
  NQT : forall t0 : nat,
        t < t0 < succn t -> ~ busy_interval.quiet_time arr_seq sched j t0
  REL : t <= job_arrival j < succn t
  QTt : busy_interval.quiet_time arr_seq sched j t
  H_busy_interval_prefix : busy_interval_prefix t (succn t)
  ============================
  hep_job j j

subgoal 2 (ID 1997) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j
subgoal 3 (ID 1996) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

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


        + by apply (H_priority_is_reflexive 0).

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

2 subgoals (ID 1997)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  CONTR : t2 < succn t1
  ============================
  exists jhp : Job,
    arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

subgoal 2 (ID 1996) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

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


      - by exfalso; move_neq_down CONTR; eapply leq_ltn_trans; eauto 2.

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

1 subgoal (ID 1996)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  ============================
  exists jhp : Job,
    arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

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


      - have EX: hp__seq: seq Job,
         j__hp, j__hp \in hp__seq arrives_in arr_seq j__hp job_pending_at j__hp t hep_job j__hp j.

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

2 subgoals (ID 2344)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  ============================
  exists hp__seq : seq Job,
    forall j__hp : Job,
    j__hp \in hp__seq <->
    arrives_in arr_seq j__hp /\ job_pending_at j__hp t /\ hep_job j__hp j

subgoal 2 (ID 2346) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

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


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

1 subgoal (ID 2344)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  ============================
  exists hp__seq : seq Job,
    forall j__hp : Job,
    j__hp \in hp__seq <->
    arrives_in arr_seq j__hp /\ job_pending_at j__hp t /\ hep_job j__hp j

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


(filter (fun jo(job_pending_at jo t) && (hep_job jo j)) (arrivals_between 0 t.+1)).

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

1 subgoal (ID 2352)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  ============================
  forall j__hp : Job,
  j__hp
    \in [seq jo <- arrivals_between 0 (succn t)
           | job_pending_at jo t
           & hep_job jo j] <->
  arrives_in arr_seq j__hp /\ job_pending_at j__hp t /\ hep_job j__hp j

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


          intros; split; intros T.

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

2 subgoals (ID 2357)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  j__hp : Job
  T : j__hp
        \in [seq jo <- arrivals_between 0 (succn t)
               | job_pending_at jo t
               & hep_job jo j]
  ============================
  arrives_in arr_seq j__hp /\ job_pending_at j__hp t /\ hep_job j__hp j

subgoal 2 (ID 2358) is:
 j__hp
   \in [seq jo <- arrivals_between 0 (succn t)
          | job_pending_at jo t
          & hep_job jo j]

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


          - move: T; rewrite mem_filter; move ⇒ /andP [/andP [PEN HP] IN].

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

2 subgoals (ID 2444)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  j__hp : Job
  PEN : job_pending_at j__hp t
  HP : hep_job j__hp j
  IN : j__hp \in arrivals_between 0 (succn t)
  ============================
  arrives_in arr_seq j__hp /\ job_pending_at j__hp t /\ hep_job j__hp j

subgoal 2 (ID 2358) is:
 j__hp
   \in [seq jo <- arrivals_between 0 (succn t)
          | job_pending_at jo t
          & hep_job jo j]

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


            repeat split; eauto using in_arrivals_implies_arrived.

(* ----------------------------------[ 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  j__hp : Job
  T : arrives_in arr_seq j__hp /\ job_pending_at j__hp t /\ hep_job j__hp j
  ============================
  j__hp
    \in [seq jo <- arrivals_between 0 (succn t)
           | job_pending_at jo t
           & hep_job jo j]

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


          - move: T ⇒ [ARR [PEN HP]].

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

1 subgoal (ID 2645)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  j__hp : Job
  ARR : arrives_in arr_seq j__hp
  PEN : job_pending_at j__hp t
  HP : hep_job j__hp j
  ============================
  j__hp
    \in [seq jo <- arrivals_between 0 (succn t)
           | job_pending_at jo t
           & hep_job jo j]

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


            rewrite mem_filter; apply/andP; split; first (apply/andP; split); try done.

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

1 subgoal (ID 2678)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  j__hp : Job
  ARR : arrives_in arr_seq j__hp
  PEN : job_pending_at j__hp t
  HP : hep_job j__hp j
  ============================
  j__hp \in arrivals_between 0 (succn t)

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


            eapply arrived_between_implies_in_arrivals; try done.

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

1 subgoal (ID 2732)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  j__hp : Job
  ARR : arrives_in arr_seq j__hp
  PEN : job_pending_at j__hp t
  HP : hep_job j__hp j
  ============================
  arrived_between j__hp 0 (succn t)

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


              by apply/andP; split; last rewrite ltnS; move: PEN ⇒ /andP [T _].

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

1 subgoal (ID 2346)

subgoal 1 (ID 2346) is:
 exists jhp : Job,
   arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

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


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

1 subgoal (ID 2346)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  EX : exists hp__seq : seq Job,
         forall j__hp : Job,
         j__hp \in hp__seq <->
         arrives_in arr_seq j__hp /\
         job_pending_at j__hp t /\ hep_job j__hp j
  ============================
  exists jhp : Job,
    arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

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


move: EX ⇒ [hp__seq SE]; case FL: (hp__seq) ⇒ [ | jhp jhps].

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

2 subgoals (ID 2936)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  hp__seq : seq Job
  SE : forall j__hp : Job,
       j__hp \in hp__seq <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t /\ hep_job j__hp j
  FL : hp__seq = [::]
  ============================
  exists jhp : Job,
    arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

subgoal 2 (ID 2981) is:
 exists jhp0 : Job,
   arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


        + subst hp__seq; exfalso.

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

2 subgoals (ID 2988)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  SE : forall j__hp : Job,
       j__hp \in [::] <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t /\ hep_job j__hp j
  ============================
  False

subgoal 2 (ID 2981) is:
 exists jhp0 : Job,
   arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


          move: GE; rewrite leq_eqVlt; move ⇒ /orP [/eqP EQ| GE].

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

3 subgoals (ID 3069)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  SE : forall j__hp : Job,
       j__hp \in [::] <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t /\ hep_job j__hp j
  EQ : t1 = t
  ============================
  False

subgoal 2 (ID 3070) is:
 False
subgoal 3 (ID 2981) is:
 exists jhp0 : Job,
   arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


          × subst t.

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

3 subgoals (ID 3078)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  LT : t1 < t2
  QTt : 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
  GT : succn t1 < t2
  SE : forall j__hp : Job,
       j__hp \in [::] <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t1 /\ hep_job j__hp j
  ============================
  False

subgoal 2 (ID 3070) is:
 False
subgoal 3 (ID 2981) is:
 exists jhp0 : Job,
   arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


            apply NQT with t1.+1; first by apply/andP; split.

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

3 subgoals (ID 3080)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  LT : t1 < t2
  QTt : 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
  GT : succn t1 < t2
  SE : forall j__hp : Job,
       j__hp \in [::] <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t1 /\ hep_job j__hp j
  ============================
  busy_interval.quiet_time arr_seq sched j (succn t1)

subgoal 2 (ID 3070) is:
 False
subgoal 3 (ID 2981) is:
 exists jhp0 : Job,
   arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


            intros jhp ARR HP ARRB; apply negbNE; apply/negP; intros NCOMP.

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

3 subgoals (ID 3135)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  LT : t1 < t2
  QTt : 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
  GT : succn t1 < t2
  SE : forall j__hp : Job,
       j__hp \in [::] <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t1 /\ hep_job j__hp j
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARRB : arrived_before jhp (succn t1)
  NCOMP : ~~ completed_by sched jhp (succn t1)
  ============================
  False

subgoal 2 (ID 3070) is:
 False
subgoal 3 (ID 2981) is:
 exists jhp0 : Job,
   arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


            move: (SE jhp) ⇒ [_ SE2].

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

3 subgoals (ID 3148)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  LT : t1 < t2
  QTt : 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
  GT : succn t1 < t2
  SE : forall j__hp : Job,
       j__hp \in [::] <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t1 /\ hep_job j__hp j
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARRB : arrived_before jhp (succn t1)
  NCOMP : ~~ completed_by sched jhp (succn t1)
  SE2 : arrives_in arr_seq jhp /\ job_pending_at jhp t1 /\ hep_job jhp j ->
        jhp \in [::]
  ============================
  False

subgoal 2 (ID 3070) is:
 False
subgoal 3 (ID 2981) is:
 exists jhp0 : Job,
   arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


            rewrite in_nil in SE2; feed SE2; [clear SE2 | by done].

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

3 subgoals (ID 3210)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  LT : t1 < t2
  QTt : 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
  GT : succn t1 < t2
  SE : forall j__hp : Job,
       j__hp \in [::] <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t1 /\ hep_job j__hp j
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARRB : arrived_before jhp (succn t1)
  NCOMP : ~~ completed_by sched jhp (succn t1)
  ============================
  arrives_in arr_seq jhp /\ job_pending_at jhp t1 /\ hep_job jhp j

subgoal 2 (ID 3070) is:
 False
subgoal 3 (ID 2981) is:
 exists jhp0 : Job,
   arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


            repeat split; try done; first apply/andP; split; try done.

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

3 subgoals (ID 3295)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  LT : t1 < t2
  QTt : 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
  GT : succn t1 < t2
  SE : forall j__hp : Job,
       j__hp \in [::] <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t1 /\ hep_job j__hp j
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARRB : arrived_before jhp (succn t1)
  NCOMP : ~~ completed_by sched jhp (succn t1)
  ============================
  ~~ completed_by sched jhp t1

subgoal 2 (ID 3070) is:
 False
subgoal 3 (ID 2981) is:
 exists jhp0 : Job,
   arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


            apply/negP; intros COMLP.

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

3 subgoals (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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  LT : t1 < t2
  QTt : 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
  GT : succn t1 < t2
  SE : forall j__hp : Job,
       j__hp \in [::] <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t1 /\ hep_job j__hp j
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARRB : arrived_before jhp (succn t1)
  NCOMP : ~~ completed_by sched jhp (succn t1)
  COMLP : completed_by sched jhp t1
  ============================
  False

subgoal 2 (ID 3070) is:
 False
subgoal 3 (ID 2981) is:
 exists jhp0 : Job,
   arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


            move: NCOMP ⇒ /negP NCOMP; apply: NCOMP.

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

3 subgoals (ID 3373)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  LT : t1 < t2
  QTt : 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
  GT : succn t1 < t2
  SE : forall j__hp : Job,
       j__hp \in [::] <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t1 /\ hep_job j__hp j
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARRB : arrived_before jhp (succn t1)
  COMLP : completed_by sched jhp t1
  ============================
  completed_by sched jhp (succn t1)

subgoal 2 (ID 3070) is:
 False
subgoal 3 (ID 2981) is:
 exists jhp0 : Job,
   arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


              by apply completion_monotonic with t1.

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

2 subgoals (ID 3070)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  SE : forall j__hp : Job,
       j__hp \in [::] <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t /\ hep_job j__hp j
  GE : t1 < t
  ============================
  False

subgoal 2 (ID 2981) is:
 exists jhp0 : Job,
   arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


          × apply NQT with t; first by apply/andP; split.

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

2 subgoals (ID 3381)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  SE : forall j__hp : Job,
       j__hp \in [::] <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t /\ hep_job j__hp j
  GE : t1 < t
  ============================
  busy_interval.quiet_time arr_seq sched j t

subgoal 2 (ID 2981) is:
 exists jhp0 : Job,
   arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


            intros jhp ARR HP ARRB; apply negbNE; apply/negP; intros NCOMP.

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

2 subgoals (ID 3436)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  SE : forall j__hp : Job,
       j__hp \in [::] <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t /\ hep_job j__hp j
  GE : t1 < t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARRB : arrived_before jhp t
  NCOMP : ~~ completed_by sched jhp t
  ============================
  False

subgoal 2 (ID 2981) is:
 exists jhp0 : Job,
   arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


            move: (SE jhp) ⇒ [_ SE2].

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

2 subgoals (ID 3449)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  SE : forall j__hp : Job,
       j__hp \in [::] <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t /\ hep_job j__hp j
  GE : t1 < t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARRB : arrived_before jhp t
  NCOMP : ~~ completed_by sched jhp t
  SE2 : arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j ->
        jhp \in [::]
  ============================
  False

subgoal 2 (ID 2981) is:
 exists jhp0 : Job,
   arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


            rewrite in_nil in SE2; feed SE2; [clear SE2 | by done].

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

2 subgoals (ID 3513)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  SE : forall j__hp : Job,
       j__hp \in [::] <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t /\ hep_job j__hp j
  GE : t1 < t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  ARRB : arrived_before jhp t
  NCOMP : ~~ completed_by sched jhp t
  ============================
  arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j

subgoal 2 (ID 2981) is:
 exists jhp0 : Job,
   arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


              by repeat split; auto; apply/andP; split; first apply ltnW.

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

1 subgoal (ID 2981)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  hp__seq : seq Job
  SE : forall j__hp : Job,
       j__hp \in hp__seq <->
       arrives_in arr_seq j__hp /\ job_pending_at j__hp t /\ hep_job j__hp j
  jhp : Job
  jhps : seq Job
  FL : hp__seq = jhp :: jhps
  ============================
  exists jhp0 : Job,
    arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


        + move: (SE jhp)=> [SE1 _]; subst; clear SE.

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

1 subgoal (ID 3608)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  GE : t1 <= t
  LT : t < t2
  QTt : 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
  GT : succn t1 < t2
  jhp : Job
  jhps : seq Job
  SE1 : jhp \in jhp :: jhps ->
        arrives_in arr_seq jhp /\ job_pending_at jhp t /\ hep_job jhp j
  ============================
  exists jhp0 : Job,
    arrives_in arr_seq jhp0 /\ job_pending_at jhp0 t /\ hep_job jhp0 j

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


            by jhp; apply SE1; rewrite in_cons; apply/orP; left.

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

No more subgoals.

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


    Qed.

We prove that at any time instant [t] within [t1, t2) the processor is not idle.
    Lemma not_quiet_implies_not_idle:
       t,
        t1 t < t2
        ¬ is_idle sched t.

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

1 subgoal (ID 1902)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  ============================
  forall t : nat, t1 <= t < t2 -> ~ is_idle sched t

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


    Proof.
      intros t NEQ IDLE.

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

1 subgoal (ID 1906)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  NEQ : t1 <= t < t2
  IDLE : is_idle sched t
  ============================
  False

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


      move: (pending_hp_job_exists _ NEQ) ⇒ [jhp [ARR [PEND HP]]].

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

1 subgoal (ID 1940)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  NEQ : t1 <= t < t2
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  PEND : job_pending_at jhp t
  HP : hep_job jhp j
  ============================
  False

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


      unfold work_conserving in ×.

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

1 subgoal (ID 1942)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : forall (j : Job) (t : instant),
                      arrives_in arr_seq j ->
                      backlogged sched j t ->
                      exists j_other : Job, scheduled_at sched j_other t
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  NEQ : t1 <= t < t2
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  PEND : job_pending_at jhp t
  HP : hep_job jhp j
  ============================
  False

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


      feed (H_work_conserving _ t ARR).

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

2 subgoals (ID 1945)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : forall (j : Job) (t : instant),
                      arrives_in arr_seq j ->
                      backlogged sched j t ->
                      exists j_other : Job, scheduled_at sched j_other t
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  NEQ : t1 <= t < t2
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  PEND : job_pending_at jhp t
  HP : hep_job jhp j
  ============================
  backlogged sched jhp t

subgoal 2 (ID 1950) is:
 False

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


      apply/andP; split; first by done.

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

2 subgoals (ID 1977)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : forall (j : Job) (t : instant),
                      arrives_in arr_seq j ->
                      backlogged sched j t ->
                      exists j_other : Job, scheduled_at sched j_other t
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  t : nat
  NEQ : t1 <= t < t2
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  PEND : job_pending_at jhp t
  HP : hep_job jhp j
  ============================
  ~~ scheduled_at sched jhp t

subgoal 2 (ID 1950) is:
 False

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


      move: IDLE ⇒ /eqP IDLE; rewrite scheduled_at_def IDLE; by done.

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

1 subgoal (ID 1950)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t : nat
  H_work_conserving : exists j_other : Job, scheduled_at sched j_other t
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  NEQ : t1 <= t < t2
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  PEND : job_pending_at jhp t
  HP : hep_job jhp j
  ============================
  False

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


      move: (H_work_conserving) ⇒ [jo SCHED].

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

1 subgoal (ID 2034)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t : nat
  H_work_conserving : exists j_other : Job, scheduled_at sched j_other t
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  NEQ : t1 <= t < t2
  IDLE : is_idle sched t
  jhp : Job
  ARR : arrives_in arr_seq jhp
  PEND : job_pending_at jhp t
  HP : hep_job jhp j
  jo : Job
  SCHED : scheduled_at sched jo t
  ============================
  False

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


      move: IDLE SCHED ⇒ /eqP IDLE SCHED.

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

1 subgoal (ID 2073)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t : nat
  H_work_conserving : exists j_other : Job, scheduled_at sched j_other t
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix t1 t2
  NEQ : t1 <= t < t2
  jhp : Job
  ARR : arrives_in arr_seq jhp
  PEND : job_pending_at jhp t
  HP : hep_job jhp j
  jo : Job
  IDLE : sched t = None
  SCHED : scheduled_at sched jo t
  ============================
  False

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


        by rewrite scheduled_at_def IDLE in SCHED.

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

No more subgoals.

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


    Qed.

  End ProcessorAlwaysBusy.

In section we prove a few auxiliary lemmas about quiet time and service.
Assume that the schedule is work-conserving ...
... and there are no duplicate job arrivals.
    Hypothesis H_arrival_sequence_is_a_set:
      arrival_sequence_uniq arr_seq.

Let [t1] be a quiet time.
    Variable t1 : instant.
    Hypothesis H_quiet_time : quiet_time t1.

Assume that there is no quiet time in the interval (t1, t1 + Δ].
    Variable Δ : duration.
    Hypothesis H_no_quiet_time : t, t1 < t t1 + Δ ¬ quiet_time t.

For clarity, we introduce a notion of the total service of jobs released in time interval [t_beg, t_end) during the time interval [t1, t1 + Δ).
We prove that jobs with higher-than-or-equal priority that released before time instant t1 receive no service after time instant t1.
    Lemma hep_jobs_receive_no_service_before_quiet_time:
      service_received_by_hep_jobs_released_during t1 (t1 + Δ) =
      service_received_by_hep_jobs_released_during 0 (t1 + Δ).

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

1 subgoal (ID 1895)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  service_received_by_hep_jobs_released_during t1 (t1 + Δ) =
  service_received_by_hep_jobs_released_during 0 (t1 + Δ)

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


    Proof.
      intros.
      rewrite /service_received_by_hep_jobs_released_during
              /service_of_higher_or_equal_priority_jobs
              /service_of_jobs /arrivals_between.

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

1 subgoal (ID 1911)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  \sum_(j0 <- arrival_sequence.arrivals_between arr_seq t1 (t1 + Δ) | 
  hep_job j0 j) service_during sched j0 t1 (t1 + Δ) =
  \sum_(j0 <- arrival_sequence.arrivals_between arr_seq 0 (t1 + Δ) | 
  hep_job j0 j) service_during sched j0 t1 (t1 + Δ)

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


      rewrite [in X in _ = X](arrivals_between_cat _ _ t1);
        [ | | rewrite leq_addr]; try done.

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

1 subgoal (ID 1934)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  \sum_(j0 <- arrival_sequence.arrivals_between arr_seq t1 (t1 + Δ) | 
  hep_job j0 j) service_during sched j0 t1 (t1 + Δ) =
  \sum_(j0 <- (arrival_sequence.arrivals_between arr_seq 0 t1 ++
               arrival_sequence.arrivals_between arr_seq t1 (t1 + Δ)) | 
  hep_job j0 j) service_during sched j0 t1 (t1 + Δ)

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


      rewrite big_cat //=.

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

1 subgoal (ID 1973)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  \sum_(j0 <- arrival_sequence.arrivals_between arr_seq t1 (t1 + Δ) | 
  hep_job j0 j) service_during sched j0 t1 (t1 + Δ) =
  \sum_(i <- arrival_sequence.arrivals_between arr_seq 0 t1 | 
  hep_job i j) service_during sched i t1 (t1 + Δ) +
  \sum_(i <- arrival_sequence.arrivals_between arr_seq t1 (t1 + Δ) | 
  hep_job i j) service_during sched i t1 (t1 + Δ)

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


      rewrite -{1}[\sum_(j <- arrivals_between _ (t1 + Δ) | _)
                    service_during sched j t1 (t1 + Δ)]add0n.

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

1 subgoal (ID 2011)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  0 +
  \sum_(j0 <- arrivals_between t1 (t1 + Δ) | hep_job j0 j)
     service_during sched j0 t1 (t1 + Δ) =
  \sum_(i <- arrival_sequence.arrivals_between arr_seq 0 t1 | 
  hep_job i j) service_during sched i t1 (t1 + Δ) +
  \sum_(i <- arrival_sequence.arrivals_between arr_seq t1 (t1 + Δ) | 
  hep_job i j) service_during sched i t1 (t1 + Δ)

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


      apply/eqP.
(* ----------------------------------[ coqtop ]---------------------------------

1 focused subgoal
(shelved: 1) (ID 2066)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  0 +
  \sum_(j0 <- arrivals_between t1 (t1 + Δ) | hep_job j0 j)
     service_during sched j0 t1 (t1 + Δ) ==
  \sum_(i <- arrival_sequence.arrivals_between arr_seq 0 t1 | 
  hep_job i j) service_during sched i t1 (t1 + Δ) +
  \sum_(i <- arrival_sequence.arrivals_between arr_seq t1 (t1 + Δ) | 
  hep_job i j) service_during sched i t1 (t1 + Δ)

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


rewrite eqn_add2r eq_sym exchange_big //=.

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

1 subgoal (ID 2089)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  \sum_(t1 <= j0 < t1 + Δ)
     \sum_(i <- arrival_sequence.arrivals_between arr_seq 0 t1 | 
     hep_job i j) (sched j0 == Some i) == 0

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


      rewrite big1_seq //.

(* ----------------------------------[ 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  forall i : nat_eqType,
  true && (i \in index_iota t1 (t1 + Δ)) ->
  \sum_(i0 <- arrival_sequence.arrivals_between arr_seq 0 t1 | 
  hep_job i0 j) (sched i == Some i0) = 0

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


      movet' /andP [_ NEQ]; rewrite mem_iota in NEQ.

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

1 subgoal (ID 2236)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat_eqType
  NEQ : t1 <= t' < t1 + (t1 + Δ - t1)
  ============================
  \sum_(i <- arrival_sequence.arrivals_between arr_seq 0 t1 | 
  hep_job i j) (sched t' == Some i) = 0

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


      rewrite big1_seq //.

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

1 subgoal (ID 2247)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat_eqType
  NEQ : t1 <= t' < t1 + (t1 + Δ - t1)
  ============================
  forall i : Job,
  hep_job i j && (i \in arrival_sequence.arrivals_between arr_seq 0 t1) ->
  (sched t' == Some i) = 0

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


      movejhp /andP [HP ARR].

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

1 subgoal (ID 2314)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat_eqType
  NEQ : t1 <= t' < t1 + (t1 + Δ - t1)
  jhp : Job
  HP : hep_job jhp j
  ARR : jhp \in arrival_sequence.arrivals_between arr_seq 0 t1
  ============================
  (sched t' == Some jhp) = 0

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


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

1 subgoal (ID 2372)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat_eqType
  NEQ : t1 <= t' < t1 + (t1 + Δ - t1)
  jhp : Job
  HP : hep_job jhp j
  ARR : jhp \in arrival_sequence.arrivals_between arr_seq 0 t1
  ============================
  sched t' != Some jhp

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


rewrite -scheduled_at_def.

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

1 subgoal (ID 2379)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat_eqType
  NEQ : t1 <= t' < t1 + (t1 + Δ - t1)
  jhp : Job
  HP : hep_job jhp j
  ARR : jhp \in arrival_sequence.arrivals_between arr_seq 0 t1
  ============================
  ~~ scheduled_at sched jhp t'

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


      apply (completed_implies_not_scheduled _ _ H_completed_jobs_dont_execute).

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

1 subgoal (ID 2388)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat_eqType
  NEQ : t1 <= t' < t1 + (t1 + Δ - t1)
  jhp : Job
  HP : hep_job jhp j
  ARR : jhp \in arrival_sequence.arrivals_between arr_seq 0 t1
  ============================
  completed_by sched jhp t'

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


      apply completion_monotonic with t1; [ move: NEQ ⇒ /andP [T1 _] | ]; try done.

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

1 subgoal (ID 2394)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat_eqType
  NEQ : t1 <= t' < t1 + (t1 + Δ - t1)
  jhp : Job
  HP : hep_job jhp j
  ARR : jhp \in arrival_sequence.arrivals_between arr_seq 0 t1
  ============================
  completed_by sched jhp t1

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


      apply H_quiet_time; try done.

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

2 subgoals (ID 2459)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat_eqType
  NEQ : t1 <= t' < t1 + (t1 + Δ - t1)
  jhp : Job
  HP : hep_job jhp j
  ARR : jhp \in arrival_sequence.arrivals_between arr_seq 0 t1
  ============================
  arrives_in arr_seq jhp

subgoal 2 (ID 2461) is:
 arrived_before jhp t1

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


      - by eapply in_arrivals_implies_arrived; eauto 2.

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

1 subgoal (ID 2461)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat_eqType
  NEQ : t1 <= t' < t1 + (t1 + Δ - t1)
  jhp : Job
  HP : hep_job jhp j
  ARR : jhp \in arrival_sequence.arrivals_between arr_seq 0 t1
  ============================
  arrived_before jhp t1

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


      - by eapply in_arrivals_implies_arrived_before; eauto 2.

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

No more subgoals.

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


    Qed.

Next we prove that the total service within a "non-quiet" time interval [t1, t1 + Δ) is exactly Δ.
    Lemma no_idle_time_within_non_quiet_time_interval:
      service_of_jobs sched predT (arrivals_between 0 (t1 + Δ)) t1 (t1 + Δ) = Δ.

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

1 subgoal (ID 1902)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  service_of_jobs sched predT (arrivals_between 0 (t1 + Δ)) t1 (t1 + Δ) = Δ

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


    Proof.
      intros; unfold service_of_jobs, service_of_higher_or_equal_priority_jobs.

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

1 subgoal (ID 1904)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  \sum_(j0 <- arrivals_between 0 (t1 + Δ) | predT j0)
     service_during sched j0 t1 (t1 + Δ) = Δ

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


      rewrite -{3}[Δ](sum_of_ones t1) exchange_big //=.

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

1 subgoal (ID 1920)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  \sum_(t1 <= j0 < t1 + Δ)
     \sum_(i <- arrivals_between 0 (t1 + Δ)) (sched j0 == Some i) =
  \sum_(t1 <= x < t1 + Δ) 1

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


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

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

2 subgoals (ID 2027)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  \sum_(t1 <= j0 < t1 + Δ)
     \sum_(i <- arrivals_between 0 (t1 + Δ)) (sched j0 == Some i) <=
  \sum_(t1 <= x < t1 + Δ) 1

subgoal 2 (ID 2028) is:
 \sum_(t1 <= x < t1 + Δ) 1 <=
 \sum_(t1 <= j0 < t1 + Δ)
    \sum_(i <- arrivals_between 0 (t1 + Δ)) (sched j0 == Some i)

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


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

1 subgoal (ID 2027)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  \sum_(t1 <= j0 < t1 + Δ)
     \sum_(i <- arrivals_between 0 (t1 + Δ)) (sched j0 == Some i) <=
  \sum_(t1 <= x < t1 + Δ) 1

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


rewrite leq_sum //.

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

1 subgoal (ID 2037)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  forall i : nat,
  true -> \sum_(i0 <- arrivals_between 0 (t1 + Δ)) (sched i == Some i0) <= 1

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


        movet' _.

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

1 subgoal (ID 2066)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  ============================
  \sum_(i <- arrivals_between 0 (t1 + Δ)) (sched t' == Some i) <= 1

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


        have SCH := service_of_jobs_le_1 sched predT (arrivals_between 0 (t1 + Δ)) _ t'.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2075)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  SCH : uniq (arrivals_between 0 (t1 + Δ)) ->
        \sum_(j <- arrivals_between 0 (t1 + Δ) | predT j)
           service_at sched j t' <= 1
  ============================
  \sum_(i <- arrivals_between 0 (t1 + Δ)) (sched t' == Some i) <= 1

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


          by eauto using arrivals_uniq.

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

1 subgoal (ID 2028)

subgoal 1 (ID 2028) is:
 \sum_(t1 <= x < t1 + Δ) 1 <=
 \sum_(t1 <= j0 < t1 + Δ)
    \sum_(i <- arrivals_between 0 (t1 + Δ)) (sched j0 == Some i)

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


      }

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

1 subgoal (ID 2028)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  \sum_(t1 <= x < t1 + Δ) 1 <=
  \sum_(t1 <= j0 < t1 + Δ)
     \sum_(i <- arrivals_between 0 (t1 + Δ)) (sched j0 == Some i)

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


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

1 subgoal (ID 2028)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  \sum_(t1 <= x < t1 + Δ) 1 <=
  \sum_(t1 <= j0 < t1 + Δ)
     \sum_(i <- arrivals_between 0 (t1 + Δ)) (sched j0 == Some i)

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


rewrite [in X in X _]big_nat_cond [in X in _ X]big_nat_cond //=; rewrite leq_sum //.

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

1 subgoal (ID 2254)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  ============================
  forall i : nat,
  (t1 <= i < t1 + Δ) && true ->
  0 < \sum_(i0 <- arrivals_between 0 (t1 + Δ)) (sched i == Some i0)

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


        movet' /andP [/andP [LT GT] _].

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

1 subgoal (ID 2359)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  LT : t1 <= t'
  GT : t' < t1 + Δ
  ============================
  0 < \sum_(i <- arrivals_between 0 (t1 + Δ)) (sched t' == Some i)

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


        apply/sum_seq_gt0P.

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

1 focused subgoal
(shelved: 1) (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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  LT : t1 <= t'
  GT : t' < t1 + Δ
  ============================
  exists i : Job,
    i \in arrivals_between 0 (t1 + Δ) /\ 0 < (sched t' == Some i)

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


        ideal_proc_model_sched_case_analysis_eq sched t' jo.

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

2 subgoals (ID 2434)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  LT : t1 <= t'
  GT : t' < t1 + Δ
  Idle : is_idle sched t'
  EqIdle : sched t' = None
  ============================
  exists i : Job, i \in arrivals_between 0 (t1 + Δ) /\ 0 < (None == Some i)

subgoal 2 (ID 2480) is:
 exists i : Job,
   i \in arrivals_between 0 (t1 + Δ) /\ 0 < (sched t' == Some i)

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


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

1 subgoal (ID 2434)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  LT : t1 <= t'
  GT : t' < t1 + Δ
  Idle : is_idle sched t'
  EqIdle : sched t' = None
  ============================
  exists i : Job, i \in arrivals_between 0 (t1 + Δ) /\ 0 < (None == Some i)

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


exfalso; move: LT; rewrite leq_eqVlt; move ⇒ /orP [/eqP EQ|LT].

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

2 subgoals (ID 2563)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  GT : t' < t1 + Δ
  Idle : is_idle sched t'
  EqIdle : sched t' = None
  EQ : t1 = t'
  ============================
  False

subgoal 2 (ID 2564) is:
 False

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


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

1 subgoal (ID 2563)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  GT : t' < t1 + Δ
  Idle : is_idle sched t'
  EqIdle : sched t' = None
  EQ : t1 = t'
  ============================
  False

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


subst t'.

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

1 subgoal (ID 2574)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  EqIdle : sched t1 = None
  Idle : is_idle sched t1
  GT : t1 < t1 + Δ
  ============================
  False

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


            feed (H_no_quiet_time t1.+1); first by apply/andP; split.

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

1 subgoal (ID 2580)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : ~ quiet_time (succn t1)
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  EqIdle : sched t1 = None
  Idle : is_idle sched t1
  GT : t1 < t1 + Δ
  ============================
  False

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


            apply: H_no_quiet_time.

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

1 focused subgoal
(shelved: 1) (ID 2612)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : ~ quiet_time (succn t1)
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  EqIdle : sched t1 = None
  Idle : is_idle sched t1
  GT : t1 < t1 + Δ
  ============================
  quiet_time (succn t1)

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


              by apply idle_time_implies_quiet_time_at_the_next_time_instant.

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

2 subgoals (ID 2564)

subgoal 1 (ID 2564) is:
 False
subgoal 2 (ID 2480) is:
 exists i : Job,
   i \in arrivals_between 0 (t1 + Δ) /\ 0 < (sched t' == Some i)

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


          }

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

1 subgoal (ID 2564)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  GT : t' < t1 + Δ
  Idle : is_idle sched t'
  EqIdle : sched t' = None
  LT : t1 < t'
  ============================
  False

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


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

1 subgoal (ID 2564)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  GT : t' < t1 + Δ
  Idle : is_idle sched t'
  EqIdle : sched t' = None
  LT : t1 < t'
  ============================
  False

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


feed (H_no_quiet_time t'); first by apply/andP; split; last rewrite ltnW.

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

1 subgoal (ID 2620)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  t' : nat
  H_no_quiet_time : ~ quiet_time t'
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  GT : t' < t1 + Δ
  Idle : is_idle sched t'
  EqIdle : sched t' = None
  LT : t1 < t'
  ============================
  False

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


            apply: H_no_quiet_time; intros j_hp IN HP ARR.

(* ----------------------------------[ 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  t' : nat
  H_no_quiet_time : ~ quiet_time t'
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  GT : t' < t1 + Δ
  Idle : is_idle sched t'
  EqIdle : sched t' = None
  LT : t1 < t'
  j_hp : Job
  IN : arrives_in arr_seq j_hp
  HP : hep_job j_hp j
  ARR : arrived_before j_hp t'
  ============================
  completed_by sched j_hp t'

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


            apply contraT; intros NOTCOMP.

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

1 subgoal (ID 2665)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  t' : nat
  H_no_quiet_time : ~ quiet_time t'
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  GT : t' < t1 + Δ
  Idle : is_idle sched t'
  EqIdle : sched t' = None
  LT : t1 < t'
  j_hp : Job
  IN : arrives_in arr_seq j_hp
  HP : hep_job j_hp j
  ARR : arrived_before j_hp t'
  NOTCOMP : ~~ completed_by sched j_hp t'
  ============================
  false

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


            destruct (scheduled_at sched j_hp t') eqn:SCHEDhp;
              first by rewrite scheduled_at_def EqIdle in SCHEDhp.

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

1 subgoal (ID 2679)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  t' : nat
  H_no_quiet_time : ~ quiet_time t'
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  GT : t' < t1 + Δ
  Idle : is_idle sched t'
  EqIdle : sched t' = None
  LT : t1 < t'
  j_hp : Job
  IN : arrives_in arr_seq j_hp
  HP : hep_job j_hp j
  ARR : arrived_before j_hp t'
  NOTCOMP : ~~ completed_by sched j_hp t'
  SCHEDhp : scheduled_at sched j_hp t' = false
  ============================
  false

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


            apply negbT in SCHEDhp.

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

1 subgoal (ID 2740)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  t' : nat
  H_no_quiet_time : ~ quiet_time t'
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  GT : t' < t1 + Δ
  Idle : is_idle sched t'
  EqIdle : sched t' = None
  LT : t1 < t'
  j_hp : Job
  IN : arrives_in arr_seq j_hp
  HP : hep_job j_hp j
  ARR : arrived_before j_hp t'
  NOTCOMP : ~~ completed_by sched j_hp t'
  SCHEDhp : ~~ scheduled_at sched j_hp t'
  ============================
  false

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


            feed (H_work_conserving j_hp t' IN);
              first by repeat (apply/andP; split); first by apply ltnW.

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

1 subgoal (ID 2747)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t' : nat
  H_work_conserving : exists j_other : Job, scheduled_at sched j_other t'
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : ~ quiet_time t'
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  GT : t' < t1 + Δ
  Idle : is_idle sched t'
  EqIdle : sched t' = None
  LT : t1 < t'
  j_hp : Job
  IN : arrives_in arr_seq j_hp
  HP : hep_job j_hp j
  ARR : arrived_before j_hp t'
  NOTCOMP : ~~ completed_by sched j_hp t'
  SCHEDhp : ~~ scheduled_at sched j_hp t'
  ============================
  false

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


            move: H_work_conserving ⇒ [j_other SCHEDother].

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

1 subgoal (ID 3156)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  t' : nat
  H_work_conserving : exists j_other : Job, scheduled_at sched j_other t'
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : ~ quiet_time t'
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  GT : t' < t1 + Δ
  Idle : is_idle sched t'
  EqIdle : sched t' = None
  LT : t1 < t'
  j_hp : Job
  IN : arrives_in arr_seq j_hp
  HP : hep_job j_hp j
  ARR : arrived_before j_hp t'
  NOTCOMP : ~~ completed_by sched j_hp t'
  SCHEDhp : ~~ scheduled_at sched j_hp t'
  j_other : Job
  SCHEDother : scheduled_at sched j_other t'
  ============================
  false

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


              by rewrite scheduled_at_def EqIdle in SCHEDother.

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

1 subgoal (ID 2480)

subgoal 1 (ID 2480) is:
 exists i : Job,
   i \in arrivals_between 0 (t1 + Δ) /\ 0 < (sched t' == Some i)

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


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

1 subgoal (ID 2480)

subgoal 1 (ID 2480) is:
 exists i : Job,
   i \in arrivals_between 0 (t1 + Δ) /\ 0 < (sched t' == Some i)

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


             
        }

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

1 subgoal (ID 2480)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  LT : t1 <= t'
  GT : t' < t1 + Δ
  jo : Job
  Sched_jo : scheduled_at sched jo t'
  EqSched_jo : #|[pred x |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched t') x)
                      x x in F]| <> 0
  ============================
  exists i : Job,
    i \in arrivals_between 0 (t1 + Δ) /\ 0 < (sched t' == Some i)

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


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

1 subgoal (ID 2480)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  LT : t1 <= t'
  GT : t' < t1 + Δ
  jo : Job
  Sched_jo : scheduled_at sched jo t'
  EqSched_jo : #|[pred x |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched t') x)
                      x x in F]| <> 0
  ============================
  exists i : Job,
    i \in arrivals_between 0 (t1 + Δ) /\ 0 < (sched t' == Some i)

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


jo; split.

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

2 subgoals (ID 3222)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  LT : t1 <= t'
  GT : t' < t1 + Δ
  jo : Job
  Sched_jo : scheduled_at sched jo t'
  EqSched_jo : #|[pred x |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched t') x)
                      x x in F]| <> 0
  ============================
  jo \in arrivals_between 0 (t1 + Δ)

subgoal 2 (ID 3223) is:
 0 < (sched t' == Some jo)

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


          - apply arrived_between_implies_in_arrivals; try done.

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

3 subgoals (ID 3227)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  LT : t1 <= t'
  GT : t' < t1 + Δ
  jo : Job
  Sched_jo : scheduled_at sched jo t'
  EqSched_jo : #|[pred x |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched t') x)
                      x x in F]| <> 0
  ============================
  arrives_in arr_seq jo

subgoal 2 (ID 3228) is:
 arrived_between jo 0 (t1 + Δ)
subgoal 3 (ID 3223) is:
 0 < (sched t' == Some jo)

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


            apply H_jobs_come_from_arrival_sequence with t'; try done.

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

2 subgoals (ID 3228)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  LT : t1 <= t'
  GT : t' < t1 + Δ
  jo : Job
  Sched_jo : scheduled_at sched jo t'
  EqSched_jo : #|[pred x |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched t') x)
                      x x in F]| <> 0
  ============================
  arrived_between jo 0 (t1 + Δ)

subgoal 2 (ID 3223) is:
 0 < (sched t' == Some jo)

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


            apply/andP; split; first by done.

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

2 subgoals (ID 3307)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  LT : t1 <= t'
  GT : t' < t1 + Δ
  jo : Job
  Sched_jo : scheduled_at sched jo t'
  EqSched_jo : #|[pred x |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched t') x)
                      x x in F]| <> 0
  ============================
  job_arrival jo < t1 + Δ

subgoal 2 (ID 3223) is:
 0 < (sched t' == Some jo)

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


            apply H_jobs_must_arrive_to_execute in Sched_jo.

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

2 subgoals (ID 3308)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  LT : t1 <= t'
  GT : t' < t1 + Δ
  jo : Job
  Sched_jo : has_arrived jo t'
  EqSched_jo : #|[pred x |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched t') x)
                      x x in F]| <> 0
  ============================
  job_arrival jo < t1 + Δ

subgoal 2 (ID 3223) is:
 0 < (sched t' == Some jo)

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


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

1 subgoal (ID 3223)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t1 : instant
  H_quiet_time : quiet_time t1
  Δ : duration
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t
  service_received_by_hep_jobs_released_during := 
  fun t_beg t_end : instant =>
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between t_beg t_end) j t1 (t1 + Δ) : 
  instant -> instant -> nat
  t' : nat
  LT : t1 <= t'
  GT : t' < t1 + Δ
  jo : Job
  Sched_jo : scheduled_at sched jo t'
  EqSched_jo : #|[pred x |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core) (, scheduled_on jo (sched t') x)
                      x x in F]| <> 0
  ============================
  0 < (sched t' == Some jo)

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


          - by rewrite lt0b -scheduled_at_def.
(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

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


        }

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

No more subgoals.

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


      }

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

No more subgoals.

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


    Qed.

  End QuietTimeAndServiceOfJobs.

In this section, we show that the length of any busy interval is bounded, as long as there is enough supply to accommodate the workload of tasks with higher or equal priority.
  Section BoundingBusyInterval.

Assume that the schedule is work-conserving, ...
... and there are no duplicate job arrivals, ...
    Hypothesis H_arrival_sequence_is_a_set:
      arrival_sequence_uniq arr_seq.

... and the priority relation is reflexive and transitive.
Next, we recall the notion of workload of all jobs released in a given interval [t1, t2) that have higher-or-equal priority w.r.t the job j being analyzed.
With regard to the jobs with higher-or-equal priority that are released in a given interval [t1, t2), we also recall the service received by these jobs in the same interval [t1, t2).
    Let hp_service t1 t2 :=
      service_of_higher_or_equal_priority_jobs
        sched (arrivals_between t1 t2) j t1 t2.

Now we begin the proof. First, we show that the busy interval is bounded.
    Section BoundingBusyInterval.

Suppose that job j is pending at time t_busy.
      Variable t_busy : instant.
      Hypothesis H_j_is_pending : job_pending_at j t_busy.

First, we show that there must exist a busy interval prefix.
      Section LowerBound.

Since job j is pending, there is a (potentially unbounded) busy interval that starts no later than with the arrival of j.
        Lemma exists_busy_interval_prefix:
           t1,
            busy_interval_prefix t1 t_busy.+1
            t1 job_arrival j t_busy.

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

1 subgoal (ID 1909)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  ============================
  exists t1 : instant,
    busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


        Proof.
          rename H_j_is_pending into PEND, H_work_conserving into WORK.

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

1 subgoal (ID 1910)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  ============================
  exists t1 : instant,
    busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


          destruct ([ t:'I_t_busy.+1, quiet_time_dec t]) eqn:EX.

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

2 subgoals (ID 1924)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  EX : ~~
       FiniteQuant.quant0b (T:=ordinal_finType (succn t_busy))
         (fun t : 'I_(succn t_busy) =>
          FiniteQuant.ex (T:=ordinal_finType (succn t_busy))
            (, quiet_time_dec t) t) = true
  ============================
  exists t1 : instant,
    busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

subgoal 2 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


          - set last0 := \max_(t < t_busy.+1 | quiet_time_dec t) t.

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

2 subgoals (ID 1934)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  EX : ~~
       FiniteQuant.quant0b (T:=ordinal_finType (succn t_busy))
         (fun t : 'I_(succn t_busy) =>
          FiniteQuant.ex (T:=ordinal_finType (succn t_busy))
            (, quiet_time_dec t) t) = true
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  ============================
  exists t1 : instant,
    busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

subgoal 2 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


            move: EX ⇒ /existsP [t EX].

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

2 subgoals (ID 1976)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  ============================
  exists t1 : instant,
    busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

subgoal 2 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


            have PRED: quiet_time_dec last0 by apply (bigmax_pred t_busy.+1 (quiet_time_dec) t).

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

2 subgoals (ID 1980)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  ============================
  exists t1 : instant,
    busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

subgoal 2 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


            have QUIET: quiet_time last0.

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

3 subgoals (ID 1981)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  ============================
  quiet_time last0

subgoal 2 (ID 1983) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy
subgoal 3 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


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

1 subgoal (ID 1981)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  ============================
  quiet_time last0

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


intros j_hp IN HP ARR; move: PRED ⇒ /allP PRED; feed (PRED j_hp).

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

2 subgoals (ID 2025)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  j_hp : Job
  IN : arrives_in arr_seq j_hp
  HP : hep_job j_hp j
  ARR : arrived_before j_hp last0
  PRED : {in arrivals_before arr_seq last0,
           forall x : Job, hep_job x j ==> completed_by sched x last0}
  ============================
  j_hp \in arrivals_before arr_seq last0

subgoal 2 (ID 2030) is:
 completed_by sched j_hp last0

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


              - by eapply arrived_between_implies_in_arrivals; eauto.

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

1 subgoal (ID 2030)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  j_hp : Job
  IN : arrives_in arr_seq j_hp
  HP : hep_job j_hp j
  ARR : arrived_before j_hp last0
  PRED : (fun x : Job => is_true (hep_job x j ==> completed_by sched x last0))
           j_hp
  ============================
  completed_by sched j_hp last0

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


              - by rewrite HP implyTb in PRED.

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

2 subgoals (ID 1983)

subgoal 1 (ID 1983) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy
subgoal 2 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


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

2 subgoals (ID 1983)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : quiet_time last0
  ============================
  exists t1 : instant,
    busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

subgoal 2 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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



             last0.

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

2 subgoals (ID 2095)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : quiet_time last0
  ============================
  busy_interval_prefix last0 (succn t_busy) /\
  last0 <= job_arrival j <= t_busy

subgoal 2 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


            have JAIN: last0 job_arrival j t_busy.

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

3 subgoals (ID 2100)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : quiet_time last0
  ============================
  last0 <= job_arrival j <= t_busy

subgoal 2 (ID 2102) is:
 busy_interval_prefix last0 (succn t_busy) /\
 last0 <= job_arrival j <= t_busy
subgoal 3 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


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

1 subgoal (ID 2100)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : quiet_time last0
  ============================
  last0 <= job_arrival j <= t_busy

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


apply/andP; split; last by move: PEND ⇒ /andP [ARR _].

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

1 subgoal (ID 2128)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : quiet_time last0
  ============================
  last0 <= job_arrival j

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


              move_neq_up BEFORE.

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

1 subgoal (ID 2199)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : quiet_time last0
  BEFORE : job_arrival j < last0
  ============================
  False

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


              move: PEND ⇒ /andP [_ NOTCOMP].

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

1 subgoal (ID 2241)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : quiet_time last0
  BEFORE : job_arrival j < last0
  NOTCOMP : ~~ completed_by sched j t_busy
  ============================
  False

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


              feed (QUIET j H_from_arrival_sequence); first by apply (H_priority_is_reflexive 0).

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

1 subgoal (ID 2247)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : arrived_before j last0 -> completed_by sched j last0
  BEFORE : job_arrival j < last0
  NOTCOMP : ~~ completed_by sched j t_busy
  ============================
  False

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


              specialize (QUIET BEFORE).

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

1 subgoal (ID 2249)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : completed_by sched j last0
  BEFORE : job_arrival j < last0
  NOTCOMP : ~~ completed_by sched j t_busy
  ============================
  False

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


              apply completion_monotonic with (t' := t_busy) in QUIET; first by rewrite QUIET in NOTCOMP.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2256)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : completed_by sched j last0
  BEFORE : job_arrival j < last0
  NOTCOMP : ~~ completed_by sched j t_busy
  ============================
  last0 <= t_busy

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


                by apply bigmax_ltn_ord with (i0 := t).

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

2 subgoals (ID 2102)

subgoal 1 (ID 2102) is:
 busy_interval_prefix last0 (succn t_busy) /\
 last0 <= job_arrival j <= t_busy
subgoal 2 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


            }

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

2 subgoals (ID 2102)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : quiet_time last0
  JAIN : last0 <= job_arrival j <= t_busy
  ============================
  busy_interval_prefix last0 (succn t_busy) /\
  last0 <= job_arrival j <= t_busy

subgoal 2 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


            repeat split; try done.

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

3 subgoals (ID 2340)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : quiet_time last0
  JAIN : last0 <= job_arrival j <= t_busy
  ============================
  last0 < succn t_busy

subgoal 2 (ID 2352) is:
 forall t0 : nat,
 last0 < t0 < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t0
subgoal 3 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


            + by apply bigmax_ltn_ord with (i0 := t).

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

2 subgoals (ID 2352)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : quiet_time last0
  JAIN : last0 <= job_arrival j <= t_busy
  ============================
  forall t0 : nat,
  last0 < t0 < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t0

subgoal 2 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


            + movet0 /andP [GTlast LTbusy] QUIET0.

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

2 subgoals (ID 2455)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : quiet_time last0
  JAIN : last0 <= job_arrival j <= t_busy
  t0 : nat
  GTlast : last0 < t0
  LTbusy : t0 < succn t_busy
  QUIET0 : busy_interval.quiet_time arr_seq sched j t0
  ============================
  False

subgoal 2 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


              have PRED0: quiet_time_dec t0.

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

3 subgoals (ID 2456)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : quiet_time last0
  JAIN : last0 <= job_arrival j <= t_busy
  t0 : nat
  GTlast : last0 < t0
  LTbusy : t0 < succn t_busy
  QUIET0 : busy_interval.quiet_time arr_seq sched j t0
  ============================
  quiet_time_dec t0

subgoal 2 (ID 2458) is:
 False
subgoal 3 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


              apply/allP; intros j_hp ARR; apply/implyP; intros HP.

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

3 subgoals (ID 2514)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : quiet_time last0
  JAIN : last0 <= job_arrival j <= t_busy
  t0 : nat
  GTlast : last0 < t0
  LTbusy : t0 < succn t_busy
  QUIET0 : busy_interval.quiet_time arr_seq sched j t0
  j_hp : Job
  ARR : j_hp \in arrivals_before arr_seq t0
  HP : hep_job j_hp j
  ============================
  completed_by sched j_hp t0

subgoal 2 (ID 2458) is:
 False
subgoal 3 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


              apply QUIET0; eauto 2 using in_arrivals_implies_arrived, in_arrivals_implies_arrived_before.

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

2 subgoals (ID 2458)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : quiet_time last0
  JAIN : last0 <= job_arrival j <= t_busy
  t0 : nat
  GTlast : last0 < t0
  LTbusy : t0 < succn t_busy
  QUIET0 : busy_interval.quiet_time arr_seq sched j t0
  PRED0 : quiet_time_dec t0
  ============================
  False

subgoal 2 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


              move_neq_down GTlast.

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

2 subgoals (ID 2588)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  last0 := \max_(t < succn t_busy | quiet_time_dec t) t : nat
  t : ordinal_finType (succn t_busy)
  EX : quiet_time_dec t
  PRED : quiet_time_dec last0
  QUIET : quiet_time last0
  JAIN : last0 <= job_arrival j <= t_busy
  t0 : nat
  LTbusy : t0 < succn t_busy
  QUIET0 : busy_interval.quiet_time arr_seq sched j t0
  PRED0 : quiet_time_dec t0
  ============================
  t0 <= last0

subgoal 2 (ID 1925) is:
 exists t1 : instant,
   busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


                by eapply (@leq_bigmax_cond _ (fun (x: 'I_t_busy.+1) ⇒ quiet_time_dec x) (fun xx) (Ordinal LTbusy)).

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

1 subgoal (ID 1925)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  EX : ~~
       FiniteQuant.quant0b (T:=ordinal_finType (succn t_busy))
         (fun t : 'I_(succn t_busy) =>
          FiniteQuant.ex (T:=ordinal_finType (succn t_busy))
            (, quiet_time_dec t) t) = false
  ============================
  exists t1 : instant,
    busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


          - apply negbT in EX; rewrite negb_exists in EX; move: EX ⇒ /forallP ALL.

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

1 subgoal (ID 2675)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  ALL : forall x : ordinal_finType (succn t_busy), ~~ quiet_time_dec x
  ============================
  exists t1 : instant,
    busy_interval_prefix t1 (succn t_busy) /\ t1 <= job_arrival j <= t_busy

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


              0; split; last by apply/andP; split; last by move: PEND ⇒ /andP [ARR _].

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

1 subgoal (ID 2679)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  ALL : forall x : ordinal_finType (succn t_busy), ~~ quiet_time_dec x
  ============================
  busy_interval_prefix 0 (succn t_busy)

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


             repeat split; first by intros j_hp _ _ ARR; rewrite /arrived_before ltn0 in ARR.

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

2 subgoals (ID 2763)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  ALL : forall x : ordinal_finType (succn t_busy), ~~ quiet_time_dec x
  ============================
  forall t : nat,
  0 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t

subgoal 2 (ID 2764) is:
 0 <= job_arrival j < succn t_busy

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


             movet /andP [GE LT].

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

2 subgoals (ID 2887)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  ALL : forall x : ordinal_finType (succn t_busy), ~~ quiet_time_dec x
  t : nat
  GE : 0 < t
  LT : t < succn t_busy
  ============================
  ~ busy_interval.quiet_time arr_seq sched j t

subgoal 2 (ID 2764) is:
 0 <= job_arrival j < succn t_busy

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


             specialize (ALL (Ordinal LT)); move: ALL ⇒ /negP ALL.

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

2 subgoals (ID 2919)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  t : nat
  LT : t < succn t_busy
  GE : 0 < t
  ALL : ~ quiet_time_dec (Ordinal (n:=succn t_busy) (m:=t) LT)
  ============================
  ~ busy_interval.quiet_time arr_seq sched j t

subgoal 2 (ID 2764) is:
 0 <= job_arrival j < succn t_busy

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


             intros QUIET; apply ALL; simpl.

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

2 subgoals (ID 2924)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  t : nat
  LT : t < succn t_busy
  GE : 0 < t
  ALL : ~ quiet_time_dec (Ordinal (n:=succn t_busy) (m:=t) LT)
  QUIET : busy_interval.quiet_time arr_seq sched j t
  ============================
  quiet_time_dec t

subgoal 2 (ID 2764) is:
 0 <= job_arrival j < succn t_busy

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


             apply/allP; intros j_hp ARR; apply/implyP; intros HP.

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

2 subgoals (ID 2980)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  t : nat
  LT : t < succn t_busy
  GE : 0 < t
  ALL : ~ quiet_time_dec (Ordinal (n:=succn t_busy) (m:=t) LT)
  QUIET : busy_interval.quiet_time arr_seq sched j t
  j_hp : Job
  ARR : j_hp \in arrivals_before arr_seq t
  HP : hep_job j_hp j
  ============================
  completed_by sched j_hp t

subgoal 2 (ID 2764) is:
 0 <= job_arrival j < succn t_busy

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


             apply QUIET; eauto 2 using in_arrivals_implies_arrived, in_arrivals_implies_arrived_before.

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

1 subgoal (ID 2764)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  ALL : forall x : ordinal_finType (succn t_busy), ~~ quiet_time_dec x
  ============================
  0 <= job_arrival j < succn t_busy

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


             apply/andP; split; first by done.

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

1 subgoal (ID 3042)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  WORK : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  PEND : job_pending_at j t_busy
  ALL : forall x : ordinal_finType (succn t_busy), ~~ quiet_time_dec x
  ============================
  job_arrival j < succn t_busy

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


               by move: PEND ⇒ /andP [ARR _].

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

No more subgoals.

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


        Qed.

      End LowerBound.

Next we prove that, if there is a point where the requested workload is upper-bounded by the supply, then the busy interval eventually ends.
      Section UpperBound.

Consider any busy interval prefix of job j.
        Variable t1 : instant.
        Hypothesis H_is_busy_prefix : busy_interval_prefix t1 t_busy.+1.

Let priority_inversion_bound be a constant which bounds length of a priority inversion.
        Variable priority_inversion_bound: instant.
        Hypothesis H_priority_inversion_is_bounded:
          is_priority_inversion_bounded_by priority_inversion_bound.

Next, assume that for some positive delta, the sum of requested workload at time [t1 + delta] and constant priority_inversion_bound is bounded by delta (i.e., the supply).
        Variable delta : duration.
        Hypothesis H_delta_positive : delta > 0.
        Hypothesis H_workload_is_bounded :
          priority_inversion_bound + hp_workload t1 (t1 + delta) delta.

If there is a quiet time by time (t1 + delta), it trivially follows that the busy interval is bounded. Thus, let's consider first the harder case where there is no quiet time, which turns out to be impossible.
        Section CannotBeBusyForSoLong.

Assume that there is no quiet time in the interval (t1, t1 + delta].
          Hypothesis H_no_quiet_time:
             t, t1 < t t1 + delta ¬ quiet_time t.

Since the interval is always non-quiet, the processor is always busy with tasks of higher-or-equal priority or some lower priority job which was scheduled, i.e., the sum of service done by jobs with actual arrival time in [t1, t1 + delta)
and priority inversion equals delta.
          Lemma busy_interval_has_uninterrupted_service:
            delta priority_inversion_bound + hp_service t1 (t1 + delta).

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

1 subgoal (ID 1904)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  ============================
  delta <= priority_inversion_bound + hp_service t1 (t1 + delta)

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


          Proof.
            move: H_is_busy_prefix ⇒ [H_strictly_larger [H_quiet [_ EXj]]].

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

1 subgoal (ID 1937)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  ============================
  delta <= priority_inversion_bound + hp_service t1 (t1 + delta)

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


            destruct (delta priority_inversion_bound) eqn:KLEΔ.

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

2 subgoals (ID 1947)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : (delta <= priority_inversion_bound) = true
  ============================
  delta <= priority_inversion_bound + hp_service t1 (t1 + delta)

subgoal 2 (ID 1948) is:
 delta <= priority_inversion_bound + hp_service t1 (t1 + delta)

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


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

1 subgoal (ID 1947)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : (delta <= priority_inversion_bound) = true
  ============================
  delta <= priority_inversion_bound + hp_service t1 (t1 + delta)

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


by apply leq_trans with priority_inversion_bound; last rewrite leq_addr.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1948)

subgoal 1 (ID 1948) is:
 delta <= priority_inversion_bound + hp_service t1 (t1 + delta)

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


}

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

1 subgoal (ID 1948)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : (delta <= priority_inversion_bound) = false
  ============================
  delta <= priority_inversion_bound + hp_service t1 (t1 + delta)

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


            apply negbT in KLEΔ; rewrite -ltnNge in KLEΔ.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2013)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  ============================
  delta <= priority_inversion_bound + hp_service t1 (t1 + delta)

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


            apply leq_trans with (cumulative_priority_inversion sched j t1 (t1 + delta) + hp_service t1 (t1 + delta)).

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

2 subgoals (ID 2016)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  ============================
  delta <=
  cumulative_priority_inversion sched j t1 (t1 + delta) +
  hp_service t1 (t1 + delta)

subgoal 2 (ID 2017) is:
 cumulative_priority_inversion sched j t1 (t1 + delta) +
 hp_service t1 (t1 + delta) <=
 priority_inversion_bound + hp_service t1 (t1 + delta)

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


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

1 subgoal (ID 2016)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  ============================
  delta <=
  cumulative_priority_inversion sched j t1 (t1 + delta) +
  hp_service t1 (t1 + delta)

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


rewrite /hp_service hep_jobs_receive_no_service_before_quiet_time //.

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

1 subgoal (ID 2023)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  ============================
  delta <=
  cumulative_priority_inversion sched j t1 (t1 + delta) +
  service_of_higher_or_equal_priority_jobs sched
    (arrivals_between 0 (t1 + delta)) j t1 (t1 + delta)

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


              rewrite /service_of_higher_or_equal_priority_jobs /service_of_jobs sum_pred_diff.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2066)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  ============================
  delta <=
  cumulative_priority_inversion sched j t1 (t1 + delta) +
  (\sum_(r <- arrivals_between 0 (t1 + delta))
      service_during sched r t1 (t1 + delta) -
   \sum_(r <- arrivals_between 0 (t1 + delta) | ~~ hep_job r j)
      service_during sched r t1 (t1 + delta))

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


              rewrite addnBA; last first.

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

2 subgoals (ID 2073)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  ============================
  \sum_(r <- arrivals_between 0 (t1 + delta) | ~~ hep_job r j)
     service_during sched r t1 (t1 + delta) <=
  \sum_(r <- arrivals_between 0 (t1 + delta))
     service_during sched r t1 (t1 + delta)

subgoal 2 (ID 2072) is:
 delta <=
 cumulative_priority_inversion sched j t1 (t1 + delta) +
 \sum_(r <- arrivals_between 0 (t1 + delta))
    service_during sched r t1 (t1 + delta) -
 \sum_(r <- arrivals_between 0 (t1 + delta) | ~~ hep_job r j)
    service_during sched r t1 (t1 + delta)

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


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

1 subgoal (ID 2073)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  ============================
  \sum_(r <- arrivals_between 0 (t1 + delta) | ~~ hep_job r j)
     service_during sched r t1 (t1 + delta) <=
  \sum_(r <- arrivals_between 0 (t1 + delta))
     service_during sched r t1 (t1 + delta)

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


by rewrite big_mkcond //= leq_sum //; intros j' _; case (hep_job j' j).
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2072)

subgoal 1 (ID 2072) is:
 delta <=
 cumulative_priority_inversion sched j t1 (t1 + delta) +
 \sum_(r <- arrivals_between 0 (t1 + delta))
    service_during sched r t1 (t1 + delta) -
 \sum_(r <- arrivals_between 0 (t1 + delta) | ~~ hep_job r j)
    service_during sched r t1 (t1 + delta)
subgoal 2 (ID 2017) is:
 cumulative_priority_inversion sched j t1 (t1 + delta) +
 hp_service t1 (t1 + delta) <=
 priority_inversion_bound + hp_service t1 (t1 + delta)

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


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

1 subgoal (ID 2072)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  ============================
  delta <=
  cumulative_priority_inversion sched j t1 (t1 + delta) +
  \sum_(r <- arrivals_between 0 (t1 + delta))
     service_during sched r t1 (t1 + delta) -
  \sum_(r <- arrivals_between 0 (t1 + delta) | ~~ hep_job r j)
     service_during sched r t1 (t1 + delta)

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



              rewrite addnC -addnBA.

(* ----------------------------------[ 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  ============================
  delta <=
  \sum_(r <- arrivals_between 0 (t1 + delta))
     service_during sched r t1 (t1 + delta) +
  (cumulative_priority_inversion sched j t1 (t1 + delta) -
   \sum_(r <- arrivals_between 0 (t1 + delta) | ~~ hep_job r j)
      service_during sched r t1 (t1 + delta))

subgoal 2 (ID 2160) is:
 \sum_(r <- arrivals_between 0 (t1 + delta) | ~~ hep_job r j)
    service_during sched r t1 (t1 + delta) <=
 cumulative_priority_inversion sched j t1 (t1 + delta)

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


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

1 subgoal (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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  ============================
  delta <=
  \sum_(r <- arrivals_between 0 (t1 + delta))
     service_during sched r t1 (t1 + delta) +
  (cumulative_priority_inversion sched j t1 (t1 + delta) -
   \sum_(r <- arrivals_between 0 (t1 + delta) | ~~ hep_job r j)
      service_during sched r t1 (t1 + delta))

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


intros. have TT := no_idle_time_within_non_quiet_time_interval.

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

1 subgoal (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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  TT : work_conserving arr_seq sched ->
       arrival_sequence_uniq arr_seq ->
       forall (t1 : instant) (Δ : duration),
       (forall t : nat, t1 < t <= t1 + Δ -> ~ quiet_time t) ->
       service_of_jobs sched predT (arrivals_between 0 (t1 + Δ)) t1 (t1 + Δ) =
       Δ
  ============================
  delta <=
  \sum_(r <- arrivals_between 0 (t1 + delta))
     service_during sched r t1 (t1 + delta) +
  (cumulative_priority_inversion sched j t1 (t1 + delta) -
   \sum_(r <- arrivals_between 0 (t1 + delta) | ~~ hep_job r j)
      service_during sched r t1 (t1 + delta))

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


                  by unfold service_of_jobs in TT; rewrite TT // leq_addr.

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

2 subgoals (ID 2160)

subgoal 1 (ID 2160) is:
 \sum_(r <- arrivals_between 0 (t1 + delta) | ~~ hep_job r j)
    service_during sched r t1 (t1 + delta) <=
 cumulative_priority_inversion sched j t1 (t1 + delta)
subgoal 2 (ID 2017) is:
 cumulative_priority_inversion sched j t1 (t1 + delta) +
 hp_service t1 (t1 + delta) <=
 priority_inversion_bound + hp_service t1 (t1 + delta)

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


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

1 subgoal (ID 2160)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  ============================
  \sum_(r <- arrivals_between 0 (t1 + delta) | ~~ hep_job r j)
     service_during sched r t1 (t1 + delta) <=
  cumulative_priority_inversion sched j t1 (t1 + delta)

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



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

1 subgoal (ID 2160)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  ============================
  \sum_(r <- arrivals_between 0 (t1 + delta) | ~~ hep_job r j)
     service_during sched r t1 (t1 + delta) <=
  cumulative_priority_inversion sched j t1 (t1 + delta)

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


rewrite /cumulative_priority_inversion /is_priority_inversion exchange_big //=.

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

1 subgoal (ID 2221)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  ============================
  \sum_(t1 <= j0 < t1 + delta)
     \sum_(i <- arrivals_between 0 (t1 + delta) | 
     ~~ hep_job i j) (sched j0 == Some i) <=
  \sum_(t1 <= t < t1 + delta)
     match sched t with
     | Some jlp => ~~ hep_job jlp j
     | None => false
     end

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


                apply leq_sum_seq; movet II _.

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

1 subgoal (ID 2248)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  t : nat_eqType
  II : t \in index_iota t1 (t1 + delta)
  ============================
  \sum_(i <- arrivals_between 0 (t1 + delta) | ~~ hep_job i j)
     (sched t == Some i) <=
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end

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


                rewrite mem_index_iota in II; move: II ⇒ /andP [GEi LEt].

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

1 subgoal (ID 2350)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  t : nat_eqType
  GEi : t1 <= t
  LEt : t < t1 + delta
  ============================
  \sum_(i <- arrivals_between 0 (t1 + delta) | ~~ hep_job i j)
     (sched t == Some i) <=
  match sched t with
  | Some jlp => ~~ hep_job jlp j
  | None => false
  end

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


                case SCHED: (sched t) ⇒ [j1 | ]; simpl; first last.

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

2 subgoals (ID 2467)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  t : nat_eqType
  GEi : t1 <= t
  LEt : t < t1 + delta
  SCHED : sched t = None
  ============================
  \sum_(i <- arrivals_between 0 (t1 + delta) | ~~ hep_job i j) 0 <= 0

subgoal 2 (ID 2465) is:
 \sum_(i <- arrivals_between 0 (t1 + delta) | ~~ hep_job i j)
    (Some j1 == Some i) <= ~~ hep_job j1 j

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


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

1 subgoal (ID 2467)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  t : nat_eqType
  GEi : t1 <= t
  LEt : t < t1 + delta
  SCHED : sched t = None
  ============================
  \sum_(i <- arrivals_between 0 (t1 + delta) | ~~ hep_job i j) 0 <= 0

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


rewrite leqn0 big1_seq //.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2465)

subgoal 1 (ID 2465) is:
 \sum_(i <- arrivals_between 0 (t1 + delta) | ~~ hep_job i j)
    (Some j1 == Some i) <= ~~ hep_job j1 j
subgoal 2 (ID 2017) is:
 cumulative_priority_inversion sched j t1 (t1 + delta) +
 hp_service t1 (t1 + delta) <=
 priority_inversion_bound + hp_service t1 (t1 + delta)

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


}

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

1 subgoal (ID 2465)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  t : nat_eqType
  GEi : t1 <= t
  LEt : t < t1 + delta
  j1 : Job
  SCHED : sched t = Some j1
  ============================
  \sum_(i <- arrivals_between 0 (t1 + delta) | ~~ hep_job i j)
     (Some j1 == Some i) <= ~~ hep_job j1 j

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


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

1 subgoal (ID 2465)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  t : nat_eqType
  GEi : t1 <= t
  LEt : t < t1 + delta
  j1 : Job
  SCHED : sched t = Some j1
  ============================
  \sum_(i <- arrivals_between 0 (t1 + delta) | ~~ hep_job i j)
     (Some j1 == Some i) <= ~~ hep_job j1 j

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


case PRIO1: (hep_job j1 j); simpl; first last.

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

2 subgoals (ID 2603)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  t : nat_eqType
  GEi : t1 <= t
  LEt : t < t1 + delta
  j1 : Job
  SCHED : sched t = Some j1
  PRIO1 : hep_job j1 j = false
  ============================
  \sum_(i <- arrivals_between 0 (t1 + delta) | ~~ hep_job i j)
     (Some j1 == Some i) <= 1

subgoal 2 (ID 2601) is:
 \sum_(i <- arrivals_between 0 (t1 + delta) | ~~ hep_job i j)
    (Some j1 == Some i) <= 0

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


                  - rewrite <- SCHED.

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

2 subgoals (ID 2604)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  t : nat_eqType
  GEi : t1 <= t
  LEt : t < t1 + delta
  j1 : Job
  SCHED : sched t = Some j1
  PRIO1 : hep_job j1 j = false
  ============================
  \sum_(i <- arrivals_between 0 (t1 + delta) | ~~ hep_job i j)
     (sched t == Some i) <= 1

subgoal 2 (ID 2601) is:
 \sum_(i <- arrivals_between 0 (t1 + delta) | ~~ hep_job i j)
    (Some j1 == Some i) <= 0

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


                    have SCH := service_of_jobs_le_1 sched _ _ _ t; eauto using arrivals_uniq.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2601)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  t : nat_eqType
  GEi : t1 <= t
  LEt : t < t1 + delta
  j1 : Job
  SCHED : sched t = Some j1
  PRIO1 : hep_job j1 j = true
  ============================
  \sum_(i <- arrivals_between 0 (t1 + delta) | ~~ hep_job i j)
     (Some j1 == Some i) <= 0

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


                  - rewrite leqn0 big1_seq; first by done.

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

1 subgoal (ID 2749)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  t : nat_eqType
  GEi : t1 <= t
  LEt : t < t1 + delta
  j1 : Job
  SCHED : sched t = Some j1
  PRIO1 : hep_job j1 j = true
  ============================
  forall i : Job,
  ~~ hep_job i j && (i \in arrivals_between 0 (t1 + delta)) ->
  (Some j1 == Some i) = 0

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


                    movej2 /andP [PRIO2 ARRj2].

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

1 subgoal (ID 2790)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  t : nat_eqType
  GEi : t1 <= t
  LEt : t < t1 + delta
  j1 : Job
  SCHED : sched t = Some j1
  PRIO1 : hep_job j1 j = true
  j2 : Job
  PRIO2 : ~~ hep_job j2 j
  ARRj2 : j2 \in arrivals_between 0 (t1 + delta)
  ============================
  (Some j1 == Some j2) = 0

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


                    case EQ: (j1 == j2).

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

2 subgoals (ID 2856)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  t : nat_eqType
  GEi : t1 <= t
  LEt : t < t1 + delta
  j1 : Job
  SCHED : sched t = Some j1
  PRIO1 : hep_job j1 j = true
  j2 : Job
  PRIO2 : ~~ hep_job j2 j
  ARRj2 : j2 \in arrivals_between 0 (t1 + delta)
  EQ : (j1 == j2) = true
  ============================
  (Some j1 == Some j2) = 0

subgoal 2 (ID 2913) is:
 (Some j1 == Some j2) = 0

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


                    + by move: EQ ⇒ /eqP EQ; subst j2; rewrite PRIO1 in PRIO2.

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

1 subgoal (ID 2913)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  t : nat_eqType
  GEi : t1 <= t
  LEt : t < t1 + delta
  j1 : Job
  SCHED : sched t = Some j1
  PRIO1 : hep_job j1 j = true
  j2 : Job
  PRIO2 : ~~ hep_job j2 j
  ARRj2 : j2 \in arrivals_between 0 (t1 + delta)
  EQ : (j1 == j2) = false
  ============================
  (Some j1 == Some j2) = 0

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


                    + apply/eqP; rewrite eqb0; apply/negP; intros CONTR; move: CONTR ⇒ /eqP CONTR.

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

1 subgoal (ID 3158)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  t : nat_eqType
  GEi : t1 <= t
  LEt : t < t1 + delta
  j1 : Job
  SCHED : sched t = Some j1
  PRIO1 : hep_job j1 j = true
  j2 : Job
  PRIO2 : ~~ hep_job j2 j
  ARRj2 : j2 \in arrivals_between 0 (t1 + delta)
  EQ : (j1 == j2) = false
  CONTR : Some j1 = Some j2
  ============================
  False

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


                        by inversion CONTR; clear CONTR; subst j2; rewrite PRIO1 in PRIO2.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2017)

subgoal 1 (ID 2017) is:
 cumulative_priority_inversion sched j t1 (t1 + delta) +
 hp_service t1 (t1 + delta) <=
 priority_inversion_bound + hp_service t1 (t1 + delta)

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


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

1 subgoal (ID 2017)

subgoal 1 (ID 2017) is:
 cumulative_priority_inversion sched j t1 (t1 + delta) +
 hp_service t1 (t1 + delta) <=
 priority_inversion_bound + hp_service t1 (t1 + delta)

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


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

1 subgoal (ID 2017)

subgoal 1 (ID 2017) is:
 cumulative_priority_inversion sched j t1 (t1 + delta) +
 hp_service t1 (t1 + delta) <=
 priority_inversion_bound + hp_service t1 (t1 + delta)

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


}

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

1 subgoal (ID 2017)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  ============================
  cumulative_priority_inversion sched j t1 (t1 + delta) +
  hp_service t1 (t1 + delta) <=
  priority_inversion_bound + hp_service t1 (t1 + delta)

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


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

1 subgoal (ID 2017)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  ============================
  cumulative_priority_inversion sched j t1 (t1 + delta) +
  hp_service t1 (t1 + delta) <=
  priority_inversion_bound + hp_service t1 (t1 + delta)

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


rewrite leq_add2r.

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

1 subgoal (ID 3282)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  ============================
  cumulative_priority_inversion sched j t1 (t1 + delta) <=
  priority_inversion_bound

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


              destruct (t1 + delta t_busy.+1) eqn:NEQ; [ | apply negbT in NEQ; rewrite -ltnNge in NEQ].

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

2 subgoals (ID 3292)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  NEQ : (t1 + delta <= succn t_busy) = true
  ============================
  cumulative_priority_inversion sched j t1 (t1 + delta) <=
  priority_inversion_bound

subgoal 2 (ID 3353) is:
 cumulative_priority_inversion sched j t1 (t1 + delta) <=
 priority_inversion_bound

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


              - apply leq_trans with (cumulative_priority_inversion sched j t1 t_busy.+1); last eauto 2.

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

2 subgoals (ID 3356)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  NEQ : (t1 + delta <= succn t_busy) = true
  ============================
  cumulative_priority_inversion sched j t1 (t1 + delta) <=
  cumulative_priority_inversion sched j t1 (succn t_busy)

subgoal 2 (ID 3353) is:
 cumulative_priority_inversion sched j t1 (t1 + delta) <=
 priority_inversion_bound

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


                  by rewrite [X in _ X](@big_cat_nat _ _ _ (t1 + delta)) //=; rewrite leq_addr.

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

1 subgoal (ID 3353)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  NEQ : succn t_busy < t1 + delta
  ============================
  cumulative_priority_inversion sched j t1 (t1 + delta) <=
  priority_inversion_bound

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


              - apply H_priority_inversion_is_bounded; repeat split; try done.

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

3 subgoals (ID 3446)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  NEQ : succn t_busy < t1 + delta
  ============================
  t1 < t1 + delta

subgoal 2 (ID 3458) is:
 forall t : nat,
 t1 < t < t1 + delta -> ~ busy_interval.quiet_time arr_seq sched j t
subgoal 3 (ID 3459) is:
 t1 <= job_arrival j < t1 + delta

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


                 + by rewrite -addn1 leq_add2l.

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

2 subgoals (ID 3458)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  NEQ : succn t_busy < t1 + delta
  ============================
  forall t : nat,
  t1 < t < t1 + delta -> ~ busy_interval.quiet_time arr_seq sched j t

subgoal 2 (ID 3459) is:
 t1 <= job_arrival j < t1 + delta

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


                 + movet' /andP [LT GT]; apply H_no_quiet_time.

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

2 subgoals (ID 3586)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  NEQ : succn t_busy < t1 + delta
  t' : nat
  LT : t1 < t'
  GT : t' < t1 + delta
  ============================
  t1 < t' <= t1 + delta

subgoal 2 (ID 3459) is:
 t1 <= job_arrival j < t1 + delta

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


                     by apply/andP; split; [ | rewrite ltnW ].

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

1 subgoal (ID 3459)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  EXj : t1 <= job_arrival j < succn t_busy
  KLEΔ : priority_inversion_bound < delta
  NEQ : succn t_busy < t1 + delta
  ============================
  t1 <= job_arrival j < t1 + delta

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


                 + move: EXj ⇒ /andP [T1 T2].

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

1 subgoal (ID 3661)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  H_strictly_larger : t1 < succn t_busy
  H_quiet : busy_interval.quiet_time arr_seq sched j t1
  KLEΔ : priority_inversion_bound < delta
  NEQ : succn t_busy < t1 + delta
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < succn t_busy
  ============================
  t1 <= job_arrival j < t1 + delta

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


                     by apply/andP; split; [done | apply ltn_trans with (t_busy.+1)].

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

No more subgoals.

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


            }

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

No more subgoals.

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


          Qed.

Moreover, the fact that the interval is not quiet also implies that there's more workload requested than service received.
          Lemma busy_interval_too_much_workload:
            hp_workload t1 (t1 + delta) > hp_service t1 (t1 + delta).

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

1 subgoal (ID 1905)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  ============================
  hp_service t1 (t1 + delta) < hp_workload t1 (t1 + delta)

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


          Proof.
            have PEND := not_quiet_implies_exists_pending_job.

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

1 subgoal (ID 1910)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  PEND : forall t1 t2 : instant,
         t1 <= t2 ->
         quiet_time t1 ->
         ~ quiet_time t2 ->
         exists j_hp : Job,
           arrives_in arr_seq j_hp /\
           arrived_between j_hp t1 t2 /\
           hep_job j_hp j /\ ~ job_completed_by j_hp t2
  ============================
  hp_service t1 (t1 + delta) < hp_workload t1 (t1 + delta)

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


            rename H_no_quiet_time into NOTQUIET,
            H_is_busy_prefix into PREFIX.

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

1 subgoal (ID 1911)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  PEND : forall t1 t2 : instant,
         t1 <= t2 ->
         quiet_time t1 ->
         ~ quiet_time t2 ->
         exists j_hp : Job,
           arrives_in arr_seq j_hp /\
           arrived_between j_hp t1 t2 /\
           hep_job j_hp j /\ ~ job_completed_by j_hp t2
  ============================
  hp_service t1 (t1 + delta) < hp_workload t1 (t1 + delta)

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


            set l := arrivals_between t1 (t1 + delta).

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

1 subgoal (ID 1913)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  PEND : forall t1 t2 : instant,
         t1 <= t2 ->
         quiet_time t1 ->
         ~ quiet_time t2 ->
         exists j_hp : Job,
           arrives_in arr_seq j_hp /\
           arrived_between j_hp t1 t2 /\
           hep_job j_hp j /\ ~ job_completed_by j_hp t2
  l := arrivals_between t1 (t1 + delta) : seq Job
  ============================
  hp_service t1 (t1 + delta) < hp_workload t1 (t1 + delta)

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


            set hep := hep_job.

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

1 subgoal (ID 1917)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  PEND : forall t1 t2 : instant,
         t1 <= t2 ->
         quiet_time t1 ->
         ~ quiet_time t2 ->
         exists j_hp : Job,
           arrives_in arr_seq j_hp /\
           arrived_between j_hp t1 t2 /\
           hep_job j_hp j /\ ~ job_completed_by j_hp t2
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  ============================
  hp_service t1 (t1 + delta) < hp_workload t1 (t1 + delta)

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


            unfold hp_service, service_of_higher_or_equal_priority_jobs, service_of_jobs,
            hp_workload, workload_of_higher_or_equal_priority_jobs, workload_of_jobs.

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

1 subgoal (ID 1919)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  PEND : forall t1 t2 : instant,
         t1 <= t2 ->
         quiet_time t1 ->
         ~ quiet_time t2 ->
         exists j_hp : Job,
           arrives_in arr_seq j_hp /\
           arrived_between j_hp t1 t2 /\
           hep_job j_hp j /\ ~ job_completed_by j_hp t2
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  ============================
  \sum_(j0 <- arrivals_between t1 (t1 + delta) | hep_job j0 j)
     service_during sched j0 t1 (t1 + delta) <
  \sum_(j0 <- arrivals_between t1 (t1 + delta) | hep_job j0 j) job_cost j0

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


            fold arrivals_between l hep.

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

1 subgoal (ID 1921)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  PEND : forall t1 t2 : instant,
         t1 <= t2 ->
         quiet_time t1 ->
         ~ quiet_time t2 ->
         exists j_hp : Job,
           arrives_in arr_seq j_hp /\
           arrived_between j_hp t1 t2 /\
           hep_job j_hp j /\ ~ job_completed_by j_hp t2
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  ============================
  \sum_(j0 <- l | hep j0 j) service_during sched j0 t1 (t1 + delta) <
  \sum_(j0 <- l | hep j0 j) job_cost j0

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


            move: (PREFIX) ⇒ [_ [QUIET _]].

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

1 subgoal (ID 1944)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  PEND : forall t1 t2 : instant,
         t1 <= t2 ->
         quiet_time t1 ->
         ~ quiet_time t2 ->
         exists j_hp : Job,
           arrives_in arr_seq j_hp /\
           arrived_between j_hp t1 t2 /\
           hep_job j_hp j /\ ~ job_completed_by j_hp t2
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  ============================
  \sum_(j0 <- l | hep j0 j) service_during sched j0 t1 (t1 + delta) <
  \sum_(j0 <- l | hep j0 j) job_cost j0

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


            move: (NOTQUIET) ⇒ NOTQUIET'.

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

1 subgoal (ID 1946)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  PEND : forall t1 t2 : instant,
         t1 <= t2 ->
         quiet_time t1 ->
         ~ quiet_time t2 ->
         exists j_hp : Job,
           arrives_in arr_seq j_hp /\
           arrived_between j_hp t1 t2 /\
           hep_job j_hp j /\ ~ job_completed_by j_hp t2
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  ============================
  \sum_(j0 <- l | hep j0 j) service_during sched j0 t1 (t1 + delta) <
  \sum_(j0 <- l | hep j0 j) job_cost j0

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


            feed (NOTQUIET' (t1 + delta)).

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

2 subgoals (ID 1947)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  PEND : forall t1 t2 : instant,
         t1 <= t2 ->
         quiet_time t1 ->
         ~ quiet_time t2 ->
         exists j_hp : Job,
           arrives_in arr_seq j_hp /\
           arrived_between j_hp t1 t2 /\
           hep_job j_hp j /\ ~ job_completed_by j_hp t2
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  ============================
  t1 < t1 + delta <= t1 + delta

subgoal 2 (ID 1952) is:
 \sum_(j0 <- l | hep j0 j) service_during sched j0 t1 (t1 + delta) <
 \sum_(j0 <- l | hep j0 j) job_cost j0

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


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

1 subgoal (ID 1947)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  PEND : forall t1 t2 : instant,
         t1 <= t2 ->
         quiet_time t1 ->
         ~ quiet_time t2 ->
         exists j_hp : Job,
           arrives_in arr_seq j_hp /\
           arrived_between j_hp t1 t2 /\
           hep_job j_hp j /\ ~ job_completed_by j_hp t2
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  ============================
  t1 < t1 + delta <= t1 + delta

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


by apply/andP; split; first
                by rewrite -addn1 leq_add2l.

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

1 subgoal (ID 1952)

subgoal 1 (ID 1952) is:
 \sum_(j0 <- l | hep j0 j) service_during sched j0 t1 (t1 + delta) <
 \sum_(j0 <- l | hep j0 j) job_cost j0

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


            }

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

1 subgoal (ID 1952)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  PEND : forall t1 t2 : instant,
         t1 <= t2 ->
         quiet_time t1 ->
         ~ quiet_time t2 ->
         exists j_hp : Job,
           arrives_in arr_seq j_hp /\
           arrived_between j_hp t1 t2 /\
           hep_job j_hp j /\ ~ job_completed_by j_hp t2
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  ============================
  \sum_(j0 <- l | hep j0 j) service_during sched j0 t1 (t1 + delta) <
  \sum_(j0 <- l | hep j0 j) job_cost j0

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


            feed (PEND t1 (t1 + delta)); first by apply leq_addr.

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

1 subgoal (ID 1993)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  PEND : quiet_time t1 ->
         ~ quiet_time (t1 + delta) ->
         exists j_hp : Job,
           arrives_in arr_seq j_hp /\
           arrived_between j_hp t1 (t1 + delta) /\
           hep_job j_hp j /\ ~ job_completed_by j_hp (t1 + delta)
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  ============================
  \sum_(j0 <- l | hep j0 j) service_during sched j0 t1 (t1 + delta) <
  \sum_(j0 <- l | hep j0 j) job_cost j0

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


            specialize (PEND QUIET NOTQUIET').

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

1 subgoal (ID 1995)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  PEND : exists j_hp : Job,
           arrives_in arr_seq j_hp /\
           arrived_between j_hp t1 (t1 + delta) /\
           hep_job j_hp j /\ ~ job_completed_by j_hp (t1 + delta)
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  ============================
  \sum_(j0 <- l | hep j0 j) service_during sched j0 t1 (t1 + delta) <
  \sum_(j0 <- l | hep j0 j) job_cost j0

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


            move: PEND ⇒ [j0 [ARR0 [/andP [GE0 LT0] [HP0 NOTCOMP0]]]].

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

1 subgoal (ID 2077)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  ============================
  \sum_(j1 <- l | hep j1 j) service_during sched j1 t1 (t1 + delta) <
  \sum_(j1 <- l | hep j1 j) job_cost j1

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


            have IN0: j0 \in l.

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

2 subgoals (ID 2082)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  ============================
  j0 \in l

subgoal 2 (ID 2084) is:
 \sum_(j1 <- l | hep j1 j) service_during sched j1 t1 (t1 + delta) <
 \sum_(j1 <- l | hep j1 j) job_cost j1

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


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

1 subgoal (ID 2082)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  ============================
  j0 \in l

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


by eapply arrived_between_implies_in_arrivals; eauto 1; apply/andP; split.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2084)

subgoal 1 (ID 2084) is:
 \sum_(j1 <- l | hep j1 j) service_during sched j1 t1 (t1 + delta) <
 \sum_(j1 <- l | hep j1 j) job_cost j1

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


}

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

1 subgoal (ID 2084)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  IN0 : j0 \in l
  ============================
  \sum_(j1 <- l | hep j1 j) service_during sched j1 t1 (t1 + delta) <
  \sum_(j1 <- l | hep j1 j) job_cost j1

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


            have UNIQ: uniq l by eapply arrivals_uniq; eauto 1.

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

1 subgoal (ID 2128)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  IN0 : j0 \in l
  UNIQ : uniq l
  ============================
  \sum_(j1 <- l | hep j1 j) service_during sched j1 t1 (t1 + delta) <
  \sum_(j1 <- l | hep j1 j) job_cost j1

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


            rewrite big_mkcond [\sum_(_ <- _ | _ _ _)_]big_mkcond //=.

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

1 subgoal (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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  IN0 : j0 \in l
  UNIQ : uniq l
  ============================
  \sum_(i <- l)
     (if hep i j then service_during sched i t1 (t1 + delta) else 0) <
  \sum_(i <- l) (if hep i j then job_cost i else 0)

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


            rewrite (bigD1_seq j0); [simpl | by done | by done].

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

1 subgoal (ID 2208)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  IN0 : j0 \in l
  UNIQ : uniq l
  ============================
  (if hep j0 j then service_during sched j0 t1 (t1 + delta) else 0) +
  \sum_(i <- l | i != j0)
     (if hep i j then service_during sched i t1 (t1 + delta) else 0) <
  \sum_(i <- l) (if hep i j then job_cost i else 0)

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


            rewrite (bigD1_seq j0); [simpl | by done | by done].

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

1 subgoal (ID 2227)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  IN0 : j0 \in l
  UNIQ : uniq l
  ============================
  (if hep j0 j then service_during sched j0 t1 (t1 + delta) else 0) +
  \sum_(i <- l | i != j0)
     (if hep i j then service_during sched i t1 (t1 + delta) else 0) <
  (if hep j0 j then job_cost j0 else 0) +
  \sum_(i <- l | i != j0) (if hep i j then job_cost i else 0)

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


            rewrite /hep HP0.

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

1 subgoal (ID 2230)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  IN0 : j0 \in l
  UNIQ : uniq l
  ============================
  service_during sched j0 t1 (t1 + delta) +
  \sum_(i <- l | i != j0)
     (if hep_job i j then service_during sched i t1 (t1 + delta) else 0) <
  job_cost j0 +
  \sum_(i <- l | i != j0) (if hep_job i j then job_cost i else 0)

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


            rewrite -add1n addnA [1 + _]addnC addn1.

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

1 subgoal (ID 2251)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  IN0 : j0 \in l
  UNIQ : uniq l
  ============================
  succn (service_during sched j0 t1 (t1 + delta)) +
  \sum_(i <- l | i != j0)
     (if hep_job i j then service_during sched i t1 (t1 + delta) else 0) <=
  job_cost j0 +
  \sum_(i <- l | i != j0) (if hep_job i j then job_cost i else 0)

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


            apply leq_add; last first.

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

2 subgoals (ID 2253)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  IN0 : j0 \in l
  UNIQ : uniq l
  ============================
  \sum_(i <- l | i != j0)
     (if hep_job i j then service_during sched i t1 (t1 + delta) else 0) <=
  \sum_(i <- l | i != j0) (if hep_job i j then job_cost i else 0)

subgoal 2 (ID 2252) is:
 service_during sched j0 t1 (t1 + delta) < job_cost j0

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


            {

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

1 subgoal (ID 2253)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  IN0 : j0 \in l
  UNIQ : uniq l
  ============================
  \sum_(i <- l | i != j0)
     (if hep_job i j then service_during sched i t1 (t1 + delta) else 0) <=
  \sum_(i <- l | i != j0) (if hep_job i j then job_cost i else 0)

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


              apply leq_sum; intros j1 NEQ.

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

1 subgoal (ID 2256)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  IN0 : j0 \in l
  UNIQ : uniq l
  j1 : Job
  NEQ : j1 != j0
  ============================
  (if hep_job j1 j then service_during sched j1 t1 (t1 + delta) else 0) <=
  (if hep_job j1 j then job_cost j1 else 0)

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


              destruct (hep_job j1 j); last by done.

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

1 subgoal (ID 2265)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  IN0 : j0 \in l
  UNIQ : uniq l
  j1 : Job
  NEQ : j1 != j0
  ============================
  service_during sched j1 t1 (t1 + delta) <= job_cost j1

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


                by apply cumulative_service_le_job_cost, ideal_proc_model_provides_unit_service.

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

1 subgoal (ID 2252)

subgoal 1 (ID 2252) is:
 service_during sched j0 t1 (t1 + delta) < job_cost j0

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


            }

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

1 subgoal (ID 2252)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  IN0 : j0 \in l
  UNIQ : uniq l
  ============================
  service_during sched j0 t1 (t1 + delta) < job_cost j0

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


            rewrite ignore_service_before_arrival; rewrite //; [| by apply ltnW].

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

1 subgoal (ID 2289)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  IN0 : j0 \in l
  UNIQ : uniq l
  ============================
  service_during sched j0 (job_arrival j0) (t1 + delta) < job_cost j0

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


            rewrite <- ignore_service_before_arrival with (t2:=0); rewrite //; [|by apply ltnW].

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

1 subgoal (ID 2345)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  NOTQUIET : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  l := arrivals_between t1 (t1 + delta) : seq Job
  hep := hep_job : rel Job
  QUIET : busy_interval.quiet_time arr_seq sched j t1
  NOTQUIET' : ~ quiet_time (t1 + delta)
  j0 : Job
  ARR0 : arrives_in arr_seq j0
  GE0 : t1 <= job_arrival j0
  LT0 : job_arrival j0 < t1 + delta
  HP0 : hep_job j0 j
  NOTCOMP0 : ~ job_completed_by j0 (t1 + delta)
  IN0 : j0 \in l
  UNIQ : uniq l
  ============================
  service_during sched j0 0 (t1 + delta) < job_cost j0

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


              by rewrite ltnNge; apply/negP.

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

No more subgoals.

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


          Qed.

Using the two lemmas above, we infer that the workload is larger than the interval length. However, this contradicts the assumption H_workload_is_bounded.
          Corollary busy_interval_workload_larger_than_interval:
            priority_inversion_bound + hp_workload t1 (t1 + delta) > delta.

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

1 subgoal (ID 1906)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  ============================
  delta < priority_inversion_bound + hp_workload t1 (t1 + delta)

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


          Proof.
            apply leq_ltn_trans with (priority_inversion_bound + hp_service t1 (t1 + delta)).

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

2 subgoals (ID 1907)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  ============================
  delta <= priority_inversion_bound + hp_service t1 (t1 + delta)

subgoal 2 (ID 1908) is:
 priority_inversion_bound + hp_service t1 (t1 + delta) <
 priority_inversion_bound + hp_workload t1 (t1 + delta)

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


            apply busy_interval_has_uninterrupted_service.

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

1 subgoal (ID 1908)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  ============================
  priority_inversion_bound + hp_service t1 (t1 + delta) <
  priority_inversion_bound + hp_workload t1 (t1 + delta)

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


            rewrite ltn_add2l.

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

1 subgoal (ID 1913)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  H_no_quiet_time : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  ============================
  hp_service t1 (t1 + delta) < hp_workload t1 (t1 + delta)

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


            apply busy_interval_too_much_workload.

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

No more subgoals.

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


          Qed.

        End CannotBeBusyForSoLong.

Since the interval cannot remain busy for so long, we prove that the busy interval finishes at some point t2 <= t1 + delta.
        Lemma busy_interval_is_bounded:
           t2,
            t2 t1 + delta
            busy_interval t1 t2.

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

1 subgoal (ID 1905)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  ============================
  exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


        Proof.
          move: H_is_busy_prefix ⇒ [LT [QT [NQ NEQ]]].

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

1 subgoal (ID 1937)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  ============================
  exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


          destruct ([ t2:'I_(t1 + delta).+1, (t2 > t1) && quiet_time_dec t2]) eqn:EX.

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

2 subgoals (ID 1951)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  EX : ~~
       FiniteQuant.quant0b (T:=ordinal_finType (succn (t1 + delta)))
         (fun t2 : 'I_(succn (t1 + delta)) =>
          FiniteQuant.ex (T:=ordinal_finType (succn (t1 + delta)))
            (, (t1 < t2) && quiet_time_dec t2) t2) = true
  ============================
  exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

subgoal 2 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


          - have EX': (t2 : instant), ((t1 < t2 t1 + delta) && quiet_time_dec t2).

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

3 subgoals (ID 1954)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  EX : ~~
       FiniteQuant.quant0b (T:=ordinal_finType (succn (t1 + delta)))
         (fun t2 : 'I_(succn (t1 + delta)) =>
          FiniteQuant.ex (T:=ordinal_finType (succn (t1 + delta)))
            (, (t1 < t2) && quiet_time_dec t2) t2) = true
  ============================
  exists t2 : instant, (t1 < t2 <= t1 + delta) && quiet_time_dec t2

subgoal 2 (ID 1956) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2
subgoal 3 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


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

1 subgoal (ID 1954)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  EX : ~~
       FiniteQuant.quant0b (T:=ordinal_finType (succn (t1 + delta)))
         (fun t2 : 'I_(succn (t1 + delta)) =>
          FiniteQuant.ex (T:=ordinal_finType (succn (t1 + delta)))
            (, (t1 < t2) && quiet_time_dec t2) t2) = true
  ============================
  exists t2 : instant, (t1 < t2 <= t1 + delta) && quiet_time_dec t2

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


move: EX ⇒ /existsP [t2 /andP [LE QUIET]].

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

1 subgoal (ID 2036)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t2 : ordinal_finType (succn (t1 + delta))
  LE : t1 < t2
  QUIET : quiet_time_dec t2
  ============================
  exists t0 : instant, (t1 < t0 <= t1 + delta) && quiet_time_dec t0

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


               t2; apply/andP; split; last by done.

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

1 subgoal (ID 2064)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t2 : ordinal_finType (succn (t1 + delta))
  LE : t1 < t2
  QUIET : quiet_time_dec t2
  ============================
  t1 < t2 <= t1 + delta

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


                by apply/andP; split; last (rewrite -ltnS; apply ltn_ord).
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1956)

subgoal 1 (ID 1956) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2
subgoal 2 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


}

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

2 subgoals (ID 1956)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  EX : ~~
       FiniteQuant.quant0b (T:=ordinal_finType (succn (t1 + delta)))
         (fun t2 : 'I_(succn (t1 + delta)) =>
          FiniteQuant.ex (T:=ordinal_finType (succn (t1 + delta)))
            (, (t1 < t2) && quiet_time_dec t2) t2) = true
  EX' : exists t2 : instant, (t1 < t2 <= t1 + delta) && quiet_time_dec t2
  ============================
  exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

subgoal 2 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


            move: (ex_minnP EX') ⇒ [t2 /andP [/andP [GT LE] QUIET] MIN]; clear EX EX'.

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

2 subgoals (ID 2188)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  ============================
  exists t0 : nat, t0 <= t1 + delta /\ busy_interval t1 t0

subgoal 2 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


             t2; split; [ | split; [repeat split | ]]; try done.

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

4 subgoals (ID 2210)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  ============================
  forall t : nat, t1 < t < t2 -> ~ busy_interval.quiet_time arr_seq sched j t

subgoal 2 (ID 2211) is:
 t1 <= job_arrival j < t2
subgoal 3 (ID 2196) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 4 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


            + movet /andP [GT1 LT2] BUG.
(* ----------------------------------[ coqtop ]---------------------------------

4 subgoals (ID 2335)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  t : nat
  GT1 : t1 < t
  LT2 : t < t2
  BUG : busy_interval.quiet_time arr_seq sched j t
  ============================
  False

subgoal 2 (ID 2211) is:
 t1 <= job_arrival j < t2
subgoal 3 (ID 2196) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 4 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


              feed (MIN t); first (apply/andP; split).

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

6 subgoals (ID 2367)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  t : nat
  GT1 : t1 < t
  LT2 : t < t2
  BUG : busy_interval.quiet_time arr_seq sched j t
  ============================
  t1 < t <= t1 + delta

subgoal 2 (ID 2368) is:
 quiet_time_dec t
subgoal 3 (ID 2341) is:
 False
subgoal 4 (ID 2211) is:
 t1 <= job_arrival j < t2
subgoal 5 (ID 2196) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 6 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


              × by apply/andP; split; last by apply leq_trans with (n := t2); eauto using ltnW.
(* ----------------------------------[ coqtop ]---------------------------------

5 subgoals (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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  t : nat
  GT1 : t1 < t
  LT2 : t < t2
  BUG : busy_interval.quiet_time arr_seq sched j t
  ============================
  quiet_time_dec t

subgoal 2 (ID 2341) is:
 False
subgoal 3 (ID 2211) is:
 t1 <= job_arrival j < t2
subgoal 4 (ID 2196) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 5 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


              × apply/allP; intros j_hp ARR; apply/implyP; intro HP.

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

5 subgoals (ID 2466)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  t : nat
  GT1 : t1 < t
  LT2 : t < t2
  BUG : busy_interval.quiet_time arr_seq sched j t
  j_hp : Job
  ARR : j_hp \in arrivals_before arr_seq t
  HP : hep_job j_hp j
  ============================
  completed_by sched j_hp t

subgoal 2 (ID 2341) is:
 False
subgoal 3 (ID 2211) is:
 t1 <= job_arrival j < t2
subgoal 4 (ID 2196) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 5 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


                apply BUG; eauto 2 using in_arrivals_implies_arrived, ARR, in_arrivals_implies_arrived_before.

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

4 subgoals (ID 2341)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  t : nat
  MIN : t2 <= t
  GT1 : t1 < t
  LT2 : t < t2
  BUG : busy_interval.quiet_time arr_seq sched j t
  ============================
  False

subgoal 2 (ID 2211) is:
 t1 <= job_arrival j < t2
subgoal 3 (ID 2196) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 4 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


                  by apply leq_ltn_trans with (p := t2) in MIN; first by rewrite ltnn in MIN.

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

3 subgoals (ID 2211)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  ============================
  t1 <= job_arrival j < t2

subgoal 2 (ID 2196) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 3 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


            + move: NEQ ⇒ /andP [IN1 IN2].

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

3 subgoals (ID 2637)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  IN1 : t1 <= job_arrival j
  IN2 : job_arrival j < succn t_busy
  ============================
  t1 <= job_arrival j < t2

subgoal 2 (ID 2196) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 3 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


              apply/andP; split; first by done.

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

3 subgoals (ID 2664)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  IN1 : t1 <= job_arrival j
  IN2 : job_arrival j < succn t_busy
  ============================
  job_arrival j < t2

subgoal 2 (ID 2196) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 3 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


              apply leq_ltn_trans with t_busy; eauto 2.

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

3 subgoals (ID 2666)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  IN1 : t1 <= job_arrival j
  IN2 : job_arrival j < succn t_busy
  ============================
  t_busy < t2

subgoal 2 (ID 2196) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 3 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


              rewrite ltnNge; apply/negP; intros CONTR.

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

3 subgoals (ID 2707)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  IN1 : t1 <= job_arrival j
  IN2 : job_arrival j < succn t_busy
  CONTR : t2 <= t_busy
  ============================
  False

subgoal 2 (ID 2196) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 3 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


              apply NQ with t2.

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

4 subgoals (ID 2708)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  IN1 : t1 <= job_arrival j
  IN2 : job_arrival j < succn t_busy
  CONTR : t2 <= t_busy
  ============================
  t1 < t2 < succn t_busy

subgoal 2 (ID 2709) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 3 (ID 2196) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 4 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


              × by apply/andP; split; last rewrite ltnS.

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

3 subgoals (ID 2709)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  IN1 : t1 <= job_arrival j
  IN2 : job_arrival j < succn t_busy
  CONTR : t2 <= t_busy
  ============================
  busy_interval.quiet_time arr_seq sched j t2

subgoal 2 (ID 2196) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 3 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


              × intros jhp ARR HP AB.
(* ----------------------------------[ coqtop ]---------------------------------

3 subgoals (ID 2745)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  IN1 : t1 <= job_arrival j
  IN2 : job_arrival j < succn t_busy
  CONTR : t2 <= t_busy
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  AB : arrived_before jhp t2
  ============================
  completed_by sched jhp t2

subgoal 2 (ID 2196) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 3 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


                move: QUIET ⇒ /allP QUIET; feed (QUIET jhp).

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

4 subgoals (ID 2782)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  IN1 : t1 <= job_arrival j
  IN2 : job_arrival j < succn t_busy
  CONTR : t2 <= t_busy
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  AB : arrived_before jhp t2
  QUIET : {in arrivals_before arr_seq t2,
            forall x : Job, hep_job x j ==> completed_by sched x t2}
  ============================
  jhp \in arrivals_before arr_seq t2

subgoal 2 (ID 2787) is:
 completed_by sched jhp t2
subgoal 3 (ID 2196) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 4 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


                eapply arrived_between_implies_in_arrivals; eauto 2.

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

3 subgoals (ID 2787)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  IN1 : t1 <= job_arrival j
  IN2 : job_arrival j < succn t_busy
  CONTR : t2 <= t_busy
  jhp : Job
  ARR : arrives_in arr_seq jhp
  HP : hep_job jhp j
  AB : arrived_before jhp t2
  QUIET : (fun x : Job => is_true (hep_job x j ==> completed_by sched x t2))
            jhp
  ============================
  completed_by sched jhp t2

subgoal 2 (ID 2196) is:
 busy_interval.quiet_time arr_seq sched j t2
subgoal 3 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


                  by move: QUIET ⇒ /implyP QUIET; apply QUIET.

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

2 subgoals (ID 2196)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  ============================
  busy_interval.quiet_time arr_seq sched j t2

subgoal 2 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


            + intros j_hp IN HP ARR.

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

2 subgoals (ID 2834)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  QUIET : quiet_time_dec t2
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  j_hp : Job
  IN : arrives_in arr_seq j_hp
  HP : hep_job j_hp j
  ARR : arrived_before j_hp t2
  ============================
  completed_by sched j_hp t2

subgoal 2 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


              move: QUIET ⇒ /allP QUIET; feed (QUIET j_hp).

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

3 subgoals (ID 2871)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  j_hp : Job
  IN : arrives_in arr_seq j_hp
  HP : hep_job j_hp j
  ARR : arrived_before j_hp t2
  QUIET : {in arrivals_before arr_seq t2,
            forall x : Job, hep_job x j ==> completed_by sched x t2}
  ============================
  j_hp \in arrivals_before arr_seq t2

subgoal 2 (ID 2876) is:
 completed_by sched j_hp t2
subgoal 3 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


              × by eapply arrived_between_implies_in_arrivals; last apply ARR.

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

2 subgoals (ID 2876)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t2 : nat
  GT : t1 < t2
  LE : t2 <= t1 + delta
  MIN : forall n : nat, (t1 < n <= t1 + delta) && quiet_time_dec n -> t2 <= n
  j_hp : Job
  IN : arrives_in arr_seq j_hp
  HP : hep_job j_hp j
  ARR : arrived_before j_hp t2
  QUIET : (fun x : Job => is_true (hep_job x j ==> completed_by sched x t2))
            j_hp
  ============================
  completed_by sched j_hp t2

subgoal 2 (ID 1952) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


              × by move: QUIET ⇒ /implyP QUIET; apply QUIET.

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

1 subgoal (ID 1952)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  EX : ~~
       FiniteQuant.quant0b (T:=ordinal_finType (succn (t1 + delta)))
         (fun t2 : 'I_(succn (t1 + delta)) =>
          FiniteQuant.ex (T:=ordinal_finType (succn (t1 + delta)))
            (, (t1 < t2) && quiet_time_dec t2) t2) = false
  ============================
  exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


          - apply negbT in EX; rewrite negb_exists in EX; move: EX ⇒ /forallP ALL'.

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

1 subgoal (ID 3005)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  ALL' : forall x : ordinal_finType (succn (t1 + delta)),
         ~~ ((t1 < x) && quiet_time_dec x)
  ============================
  exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


            have ALL: t, t1 < t t1 + delta ¬ quiet_time t.

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

2 subgoals (ID 3007)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  ALL' : forall x : ordinal_finType (succn (t1 + delta)),
         ~~ ((t1 < x) && quiet_time_dec x)
  ============================
  forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t

subgoal 2 (ID 3009) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


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

1 subgoal (ID 3007)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  ALL' : forall x : ordinal_finType (succn (t1 + delta)),
         ~~ ((t1 < x) && quiet_time_dec x)
  ============================
  forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t

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


movet /andP [GTt LEt] QUIET; rewrite -ltnS in LEt.

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

1 subgoal (ID 3113)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  ALL' : forall x : ordinal_finType (succn (t1 + delta)),
         ~~ ((t1 < x) && quiet_time_dec x)
  t : nat
  GTt : t1 < t
  QUIET : quiet_time t
  LEt : t < succn (t1 + delta)
  ============================
  False

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


              specialize (ALL' (Ordinal LEt)); rewrite negb_and /= GTt orFb in ALL'.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 3192)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t : nat
  LEt : t < succn (t1 + delta)
  GTt : t1 < t
  QUIET : quiet_time t
  ALL' : ~~ quiet_time_dec t
  ============================
  False

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


              move: ALL' ⇒ /negP ALL'; apply ALL'; clear ALL'.

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

1 subgoal (ID 3222)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t : nat
  LEt : t < succn (t1 + delta)
  GTt : t1 < t
  QUIET : quiet_time t
  ============================
  quiet_time_dec t

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


              apply/allP; intros j_hp ARR; apply/implyP; intro HP.

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

1 subgoal (ID 3278)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  t : nat
  LEt : t < succn (t1 + delta)
  GTt : t1 < t
  QUIET : quiet_time t
  j_hp : Job
  ARR : j_hp \in arrivals_before arr_seq t
  HP : hep_job j_hp j
  ============================
  completed_by sched j_hp t

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


              apply QUIET; eauto 2 using in_arrivals_implies_arrived, ARR, in_arrivals_implies_arrived_before.

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

1 subgoal (ID 3009)

subgoal 1 (ID 3009) is:
 exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


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

1 subgoal (ID 3009)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  ALL' : forall x : ordinal_finType (succn (t1 + delta)),
         ~~ ((t1 < x) && quiet_time_dec x)
  ALL : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  ============================
  exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2

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


clear ALL'; exfalso.

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

1 subgoal (ID 3315)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  ALL : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  ============================
  False

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


            have TOOMUCH := busy_interval_workload_larger_than_interval.

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

1 subgoal (ID 3320)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  ALL : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  TOOMUCH : (forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t) ->
            delta < priority_inversion_bound + hp_workload t1 (t1 + delta)
  ============================
  False

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


            have BOUNDED := H_workload_is_bounded.

(* ----------------------------------[ 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  t_busy : instant
  H_j_is_pending : job_pending_at j t_busy
  t1 : instant
  H_is_busy_prefix : busy_interval_prefix t1 (succn t_busy)
  priority_inversion_bound : instant
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : priority_inversion_bound +
                          hp_workload t1 (t1 + delta) <= delta
  LT : t1 < succn t_busy
  QT : busy_interval.quiet_time arr_seq sched j t1
  NQ : forall t : nat,
       t1 < t < succn t_busy -> ~ busy_interval.quiet_time arr_seq sched j t
  NEQ : t1 <= job_arrival j < succn t_busy
  ALL : forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t
  TOOMUCH : (forall t : nat, t1 < t <= t1 + delta -> ~ quiet_time t) ->
            delta < priority_inversion_bound + hp_workload t1 (t1 + delta)
  BOUNDED : priority_inversion_bound + hp_workload t1 (t1 + delta) <= delta
  ============================
  False

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


              by move: (leq_trans (TOOMUCH ALL) BOUNDED); rewrite ltnn.
(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

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



        Qed.

      End UpperBound.

    End BoundingBusyInterval.

In this section, we show that from a workload bound we can infer the existence of a busy interval.
Let priority_inversion_bound be a constant that bounds the length of a priority inversion.
Assume that for some positive delta, the sum of requested workload at time (t1 + delta) and priority inversion is bounded by delta (i.e., the supply).
      Variable delta: duration.
      Hypothesis H_delta_positive: delta > 0.
      Hypothesis H_workload_is_bounded:
         t, priority_inversion_bound + hp_workload t (t + delta) delta.

Next, we assume that job j has positive cost, from which we can infer that there is a time in which j is pending.
      Hypothesis H_positive_cost: job_cost j > 0.

Therefore there must exists a busy interval [t1, t2) that contains the arrival time of j.
      Corollary exists_busy_interval:
         t1 t2,
          t1 job_arrival j < t2
          t2 t1 + delta
          busy_interval t1 t2.

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

1 subgoal (ID 1914)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  ============================
  exists t1 t2 : nat,
    t1 <= job_arrival j < t2 /\ t2 <= t1 + delta /\ busy_interval t1 t2

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


      Proof.
        have PREFIX := exists_busy_interval_prefix.

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

1 subgoal (ID 1919)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  PREFIX : forall t_busy : instant,
           job_pending_at j t_busy ->
           exists t1 : instant,
             busy_interval_prefix t1 (succn t_busy) /\
             t1 <= job_arrival j <= t_busy
  ============================
  exists t1 t2 : nat,
    t1 <= job_arrival j < t2 /\ t2 <= t1 + delta /\ busy_interval t1 t2

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


        move: (H_workload_is_bounded) ⇒ WORK.

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

1 subgoal (ID 1921)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  PREFIX : forall t_busy : instant,
           job_pending_at j t_busy ->
           exists t1 : instant,
             busy_interval_prefix t1 (succn t_busy) /\
             t1 <= job_arrival j <= t_busy
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  ============================
  exists t1 t2 : nat,
    t1 <= job_arrival j < t2 /\ t2 <= t1 + delta /\ busy_interval t1 t2

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


        feed (PREFIX (job_arrival j)).

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

2 subgoals (ID 1924)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  PREFIX : forall t_busy : instant,
           job_pending_at j t_busy ->
           exists t1 : instant,
             busy_interval_prefix t1 (succn t_busy) /\
             t1 <= job_arrival j <= t_busy
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  ============================
  job_pending_at j (job_arrival j)

subgoal 2 (ID 1929) is:
 exists t1 t2 : nat,
   t1 <= job_arrival j < t2 /\ t2 <= t1 + delta /\ busy_interval t1 t2

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


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

1 subgoal (ID 1924)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  PREFIX : forall t_busy : instant,
           job_pending_at j t_busy ->
           exists t1 : instant,
             busy_interval_prefix t1 (succn t_busy) /\
             t1 <= job_arrival j <= t_busy
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  ============================
  job_pending_at j (job_arrival j)

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


apply/andP; split; first by apply leqnn.

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

1 subgoal (ID 1956)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  PREFIX : forall t_busy : instant,
           job_pending_at j t_busy ->
           exists t1 : instant,
             busy_interval_prefix t1 (succn t_busy) /\
             t1 <= job_arrival j <= t_busy
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  ============================
  ~~ completed_by sched j (job_arrival j)

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


          rewrite /completed_by /service.

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

1 subgoal (ID 1970)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  PREFIX : forall t_busy : instant,
           job_pending_at j t_busy ->
           exists t1 : instant,
             busy_interval_prefix t1 (succn t_busy) /\
             t1 <= job_arrival j <= t_busy
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  ============================
  ~~ (job_cost j <= service_during sched j 0 (job_arrival j))

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


          rewrite ignore_service_before_arrival // /service_during.

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

1 subgoal (ID 2018)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  PREFIX : forall t_busy : instant,
           job_pending_at j t_busy ->
           exists t1 : instant,
             busy_interval_prefix t1 (succn t_busy) /\
             t1 <= job_arrival j <= t_busy
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  ============================
  ~~
  (job_cost j <=
   \sum_(job_arrival j <= t < job_arrival j) service_at sched j t)

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


          rewrite big_geq; last by apply leqnn.

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

1 subgoal (ID 2028)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  PREFIX : forall t_busy : instant,
           job_pending_at j t_busy ->
           exists t1 : instant,
             busy_interval_prefix t1 (succn t_busy) /\
             t1 <= job_arrival j <= t_busy
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  ============================
  ~~ (job_cost j <= 0)

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


            by rewrite -ltnNge.

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

1 subgoal (ID 1929)

subgoal 1 (ID 1929) is:
 exists t1 t2 : nat,
   t1 <= job_arrival j < t2 /\ t2 <= t1 + delta /\ busy_interval t1 t2

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


        }

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

1 subgoal (ID 1929)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  PREFIX : exists t1 : instant,
             busy_interval_prefix t1 (succn (job_arrival j)) /\
             t1 <= job_arrival j <= job_arrival j
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  ============================
  exists t1 t2 : nat,
    t1 <= job_arrival j < t2 /\ t2 <= t1 + delta /\ busy_interval t1 t2

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


        move: PREFIX ⇒ [t1 [PREFIX /andP [GE1 GEarr]]].

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

1 subgoal (ID 2095)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  GEarr : job_arrival j <= job_arrival j
  ============================
  exists t0 t2 : nat,
    t0 <= job_arrival j < t2 /\ t2 <= t0 + delta /\ busy_interval t0 t2

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


        have BOUNDED := busy_interval_is_bounded
                          (job_arrival j) t1 PREFIX priority_inversion_bound _ delta
                          H_delta_positive.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2104)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  GEarr : job_arrival j <= job_arrival j
  BOUNDED : is_priority_inversion_bounded_by priority_inversion_bound ->
            priority_inversion_bound + hp_workload t1 (t1 + delta) <= delta ->
            exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2
  ============================
  exists t0 t2 : nat,
    t0 <= job_arrival j < t2 /\ t2 <= t0 + delta /\ busy_interval t0 t2

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



        feed_n 2 BOUNDED; try done.

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

1 subgoal (ID 2116)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  GEarr : job_arrival j <= job_arrival j
  BOUNDED : exists t2 : nat, t2 <= t1 + delta /\ busy_interval t1 t2
  ============================
  exists t0 t2 : nat,
    t0 <= job_arrival j < t2 /\ t2 <= t0 + delta /\ busy_interval t0 t2

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


        move: BOUNDED ⇒ [t2 [GE2 BUSY]].

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

1 subgoal (ID 2166)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  GEarr : job_arrival j <= job_arrival j
  t2 : nat
  GE2 : t2 <= t1 + delta
  BUSY : busy_interval t1 t2
  ============================
  exists t0 t3 : nat,
    t0 <= job_arrival j < t3 /\ t3 <= t0 + delta /\ busy_interval t0 t3

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


         t1, t2; split.

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

2 subgoals (ID 2172)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  GEarr : job_arrival j <= job_arrival j
  t2 : nat
  GE2 : t2 <= t1 + delta
  BUSY : busy_interval t1 t2
  ============================
  t1 <= job_arrival j < t2

subgoal 2 (ID 2173) is:
 t2 <= t1 + delta /\ busy_interval t1 t2

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


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

1 subgoal (ID 2172)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  GEarr : job_arrival j <= job_arrival j
  t2 : nat
  GE2 : t2 <= t1 + delta
  BUSY : busy_interval t1 t2
  ============================
  t1 <= job_arrival j < t2

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


apply/andP; split; first by done.

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

1 subgoal (ID 2200)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  GEarr : job_arrival j <= job_arrival j
  t2 : nat
  GE2 : t2 <= t1 + delta
  BUSY : busy_interval t1 t2
  ============================
  job_arrival j < t2

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


          apply contraT; rewrite -leqNgt; intro BUG.

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

1 subgoal (ID 2206)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  GEarr : job_arrival j <= job_arrival j
  t2 : nat
  GE2 : t2 <= t1 + delta
  BUSY : busy_interval t1 t2
  BUG : t2 <= job_arrival j
  ============================
  false

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


          move: BUSY PREFIX ⇒ [[LE12 _] QUIET] [_ [_ [NOTQUIET _]]].

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

1 subgoal (ID 2263)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  t1 : instant
  GE1 : t1 <= job_arrival j
  GEarr : job_arrival j <= job_arrival j
  t2 : nat
  GE2 : t2 <= t1 + delta
  BUG : t2 <= job_arrival j
  LE12 : t1 < t2
  QUIET : busy_interval.quiet_time arr_seq sched j t2
  NOTQUIET : forall t : nat,
             t1 < t < succn (job_arrival j) ->
             ~ busy_interval.quiet_time arr_seq sched j t
  ============================
  false

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


          feed (NOTQUIET t2); first by apply/andP; split.

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

1 subgoal (ID 2269)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  t1 : instant
  GE1 : t1 <= job_arrival j
  GEarr : job_arrival j <= job_arrival j
  t2 : nat
  GE2 : t2 <= t1 + delta
  BUG : t2 <= job_arrival j
  LE12 : t1 < t2
  QUIET : busy_interval.quiet_time arr_seq sched j t2
  NOTQUIET : ~ busy_interval.quiet_time arr_seq sched j t2
  ============================
  false

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


            by exfalso; apply NOTQUIET.

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

1 subgoal (ID 2173)

subgoal 1 (ID 2173) is:
 t2 <= t1 + delta /\ busy_interval t1 t2

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


        }

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

1 subgoal (ID 2173)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  H_positive_cost : 0 < job_cost j
  WORK : forall t : instant,
         priority_inversion_bound + hp_workload t (t + delta) <= delta
  t1 : instant
  PREFIX : busy_interval_prefix t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  GEarr : job_arrival j <= job_arrival j
  t2 : nat
  GE2 : t2 <= t1 + delta
  BUSY : busy_interval t1 t2
  ============================
  t2 <= t1 + delta /\ busy_interval t1 t2

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


          by split.
(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

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


      Qed.

    End BusyIntervalFromWorkloadBound.

If we know that the workload is bounded, we can also use the busy interval to infer a response-time bound.
Let priority_inversion_bound be a constant that bounds the length of a priority inversion.
Assume that for some positive delta, the sum of requested workload at time (t1 + delta) and priority inversion is bounded by delta (i.e., the supply).
      Variable delta: duration.
      Hypothesis H_delta_positive: delta > 0.
      Hypothesis H_workload_is_bounded:
         t, priority_inversion_bound + hp_workload t (t + delta) delta.

Then, job j must complete by (job_arrival j + delta).
      Lemma busy_interval_bounds_response_time:
        job_completed_by j (job_arrival j + delta).

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

1 subgoal (ID 1906)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  ============================
  job_completed_by j (job_arrival j + delta)

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


      Proof.
        have BUSY := exists_busy_interval priority_inversion_bound _ delta.

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

1 subgoal (ID 1913)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  BUSY : is_priority_inversion_bounded_by priority_inversion_bound ->
         0 < delta ->
         (forall t : instant,
          priority_inversion_bound + hp_workload t (t + delta) <= delta) ->
         0 < job_cost j ->
         exists t1 t2 : nat,
           t1 <= job_arrival j < t2 /\
           t2 <= t1 + delta /\ busy_interval t1 t2
  ============================
  job_completed_by j (job_arrival j + delta)

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


        move: (posnP (@job_cost _ H2 j)) ⇒ [ZERO|POS].

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

2 subgoals (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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  BUSY : is_priority_inversion_bounded_by priority_inversion_bound ->
         0 < delta ->
         (forall t : instant,
          priority_inversion_bound + hp_workload t (t + delta) <= delta) ->
         0 < job_cost j ->
         exists t1 t2 : nat,
           t1 <= job_arrival j < t2 /\
           t2 <= t1 + delta /\ busy_interval t1 t2
  ZERO : job_cost j = 0
  ============================
  job_completed_by j (job_arrival j + delta)

subgoal 2 (ID 1931) is:
 job_completed_by j (job_arrival j + delta)

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


        {
(* ----------------------------------[ 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  BUSY : is_priority_inversion_bounded_by priority_inversion_bound ->
         0 < delta ->
         (forall t : instant,
          priority_inversion_bound + hp_workload t (t + delta) <= delta) ->
         0 < job_cost j ->
         exists t1 t2 : nat,
           t1 <= job_arrival j < t2 /\
           t2 <= t1 + delta /\ busy_interval t1 t2
  ZERO : job_cost j = 0
  ============================
  job_completed_by j (job_arrival j + delta)

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


by rewrite /job_completed_by /completed_by ZERO.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1931)

subgoal 1 (ID 1931) is:
 job_completed_by j (job_arrival j + delta)

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


}

(* ----------------------------------[ 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  BUSY : is_priority_inversion_bounded_by priority_inversion_bound ->
         0 < delta ->
         (forall t : instant,
          priority_inversion_bound + hp_workload t (t + delta) <= delta) ->
         0 < job_cost j ->
         exists t1 t2 : nat,
           t1 <= job_arrival j < t2 /\
           t2 <= t1 + delta /\ busy_interval t1 t2
  POS : 0 < job_cost j
  ============================
  job_completed_by j (job_arrival j + delta)

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


        feed_n 4 BUSY; try by done.

(* ----------------------------------[ 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  BUSY : exists t1 t2 : nat,
           t1 <= job_arrival j < t2 /\
           t2 <= t1 + delta /\ busy_interval t1 t2
  POS : 0 < job_cost j
  ============================
  job_completed_by j (job_arrival j + delta)

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


        move: BUSY ⇒ [t1 [t2 [/andP [GE1 LT2] [GE2 BUSY]]]].

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

1 subgoal (ID 2069)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  POS : 0 < job_cost j
  t1, t2 : nat
  GE1 : t1 <= job_arrival j
  LT2 : job_arrival j < t2
  GE2 : t2 <= t1 + delta
  BUSY : busy_interval t1 t2
  ============================
  job_completed_by j (job_arrival j + delta)

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


        apply completion_monotonic with (t := t2); try (by done);
          first by apply leq_trans with (n := t1 + delta); [| by rewrite leq_add2r].

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

1 subgoal (ID 2075)
  
  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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending 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
  tsk : Task
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_task : job_of_task tsk j
  H_job_cost_positive : job_cost_positive j
  quiet_time := [eta busy_interval.quiet_time arr_seq sched j]
   : instant -> Prop
  quiet_time_dec := [eta busy_interval.quiet_time_dec arr_seq sched j]
   : instant -> bool
  busy_interval_prefix := fun t1 : instant =>
                          [eta busy_interval.busy_interval_prefix arr_seq
                                 sched j t1] : instant -> instant -> Prop
  busy_interval := fun t1 : instant =>
                   [eta busy_interval.busy_interval arr_seq sched j t1]
   : instant -> instant -> Prop
  is_priority_inversion_bounded_by := [eta priority_inversion_of_job_is_bounded_by
                                             arr_seq sched j]
   : duration -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  hp_workload := fun t1 t2 : instant =>
                 workload_of_higher_or_equal_priority_jobs j
                   (arrivals_between t1 t2) : instant -> instant -> nat
  hp_service := fun t1 t2 : instant =>
                service_of_higher_or_equal_priority_jobs sched
                  (arrivals_between t1 t2) j t1 t2
   : instant -> instant -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : is_priority_inversion_bounded_by
                                      priority_inversion_bound
  delta : duration
  H_delta_positive : 0 < delta
  H_workload_is_bounded : forall t : instant,
                          priority_inversion_bound +
                          hp_workload t (t + delta) <= delta
  POS : 0 < job_cost j
  t1, t2 : nat
  GE1 : t1 <= job_arrival j
  LT2 : job_arrival j < t2
  GE2 : t2 <= t1 + delta
  BUSY : busy_interval t1 t2
  ============================
  completed_by sched j t2

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


        apply job_completes_within_busy_interval with (t1 := t1); try by done.

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

No more subgoals.

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


      Qed.

    End ResponseTimeBoundFromBusyInterval.

  End BoundingBusyInterval.

End ExistsBusyIntervalJLFP.