Library prosa.analysis.facts.busy_interval.carry_in


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

Welcome to Coq 8.11.2 (June 2020)

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


Require Export prosa.analysis.facts.model.workload.
Require Export prosa.analysis.definitions.carry_in.
Require Export prosa.analysis.facts.busy_interval.busy_interval.

Throughout this file, we assume ideal uniprocessor 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 No Carry-In Instant

In this section, we derive an alternative condition for the existence of a busy interval. The new condition requires the total workload (instead of the high-priority workload) generated by the task set to be bounded.
Next, we derive a sufficient condition for existence of no carry-in instant for uniprocessor for JLFP schedulers
Section ExistsNoCarryIn.

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 uniprocessor 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.
The fact that there is no carry-in at time instant t trivially implies that t is a quiet time.
  Lemma no_carry_in_implies_quiet_time :
     j t,
      no_carry_in t
      quiet_time j t.

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

1 subgoal (ID 1254)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  ============================
  forall (j : Job) (t : instant), no_carry_in t -> quiet_time j t

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


  Proof.
      by intros j t FQT j_hp ARR HP BEF; apply FQT.

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

No more subgoals.

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


  Qed.

Assume that the schedule is work-conserving, ...
... and there are no duplicate job arrivals.
We show that an idle time implies no carry in at this time instant.
  Lemma idle_instant_implies_no_carry_in_at_t :
     t,
      is_idle sched t
      no_carry_in t.

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

1 subgoal (ID 1264)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  ============================
  forall t : instant, is_idle sched t -> no_carry_in t

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


  Proof.
    intros ? IDLE j ARR HA.

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

1 subgoal (ID 1270)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  IDLE : is_idle sched t
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j t
  ============================
  completed_by sched j t

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


    apply/negPn/negP; intros NCOMPL.

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

1 subgoal (ID 1370)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  IDLE : is_idle sched t
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j t
  NCOMPL : ~~ completed_by sched j t
  ============================
  False

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


    move: IDLE ⇒ /eqP IDLE.

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

1 subgoal (ID 1406)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j t
  NCOMPL : ~~ completed_by sched j t
  IDLE : sched t = None
  ============================
  False

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


    move: (H_work_conserving j t ARR) ⇒ NIDLE.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1408)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j t
  NCOMPL : ~~ completed_by sched j t
  IDLE : sched t = None
  NIDLE : backlogged sched j t ->
          exists j_other : Job, scheduled_at sched j_other t
  ============================
  False

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


    feed NIDLE.

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

2 subgoals (ID 1409)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j t
  NCOMPL : ~~ completed_by sched j t
  IDLE : sched t = None
  NIDLE : backlogged sched j t ->
          exists j_other : Job, scheduled_at sched j_other t
  ============================
  backlogged sched j t

subgoal 2 (ID 1414) is:
 False

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


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

1 subgoal (ID 1409)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j t
  NCOMPL : ~~ completed_by sched j t
  IDLE : sched t = None
  NIDLE : backlogged sched j t ->
          exists j_other : Job, scheduled_at sched j_other t
  ============================
  backlogged sched j t

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


apply/andP; split; last first.

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

2 subgoals (ID 1441)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j t
  NCOMPL : ~~ completed_by sched j t
  IDLE : sched t = None
  NIDLE : backlogged sched j t ->
          exists j_other : Job, scheduled_at sched j_other t
  ============================
  ~~ scheduled_at sched j t

subgoal 2 (ID 1440) is:
 job_ready sched j t

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


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

1 subgoal (ID 1441)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j t
  NCOMPL : ~~ completed_by sched j t
  IDLE : sched t = None
  NIDLE : backlogged sched j t ->
          exists j_other : Job, scheduled_at sched j_other t
  ============================
  ~~ scheduled_at sched j t

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


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

2 subgoals (ID 1440)

subgoal 1 (ID 1440) is:
 job_ready sched j t
subgoal 2 (ID 1414) is:
 False

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


}

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

1 subgoal (ID 1440)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j t
  NCOMPL : ~~ completed_by sched j t
  IDLE : sched t = None
  NIDLE : backlogged sched j t ->
          exists j_other : Job, scheduled_at sched j_other t
  ============================
  job_ready sched j t

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


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

1 subgoal (ID 1440)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j t
  NCOMPL : ~~ completed_by sched j t
  IDLE : sched t = None
  NIDLE : backlogged sched j t ->
          exists j_other : Job, scheduled_at sched j_other t
  ============================
  job_ready sched j t

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


by apply/andP; split; [apply ltnW | done].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1414)

subgoal 1 (ID 1414) is:
 False

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


}

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

1 subgoal (ID 1414)

subgoal 1 (ID 1414) is:
 False

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


    }

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

1 subgoal (ID 1414)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j t
  NCOMPL : ~~ completed_by sched j t
  IDLE : sched t = None
  NIDLE : exists j_other : Job, scheduled_at sched j_other t
  ============================
  False

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


    move: NIDLE ⇒ [j' SCHED].

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

1 subgoal (ID 1491)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j t
  NCOMPL : ~~ completed_by sched j t
  IDLE : sched t = None
  j' : Job
  SCHED : scheduled_at sched j' t
  ============================
  False

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


      by rewrite scheduled_at_def IDLE in SCHED.

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

No more subgoals.

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


  Qed.

Moreover, an idle time implies no carry in at the next time instant.
  Lemma idle_instant_implies_no_carry_in_at_t_pl_1 :
     t,
      is_idle sched t
      no_carry_in t.+1.

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

1 subgoal (ID 1267)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  ============================
  forall t : instant, is_idle sched t -> no_carry_in (succn t)

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


  Proof.
    intros ? IDLE j ARR HA.

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

1 subgoal (ID 1273)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  IDLE : is_idle sched t
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j (succn t)
  ============================
  completed_by sched j (succn t)

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


    apply/negPn/negP; intros NCOMPL.

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

1 subgoal (ID 1373)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  IDLE : is_idle sched t
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j (succn t)
  NCOMPL : ~~ completed_by sched j (succn t)
  ============================
  False

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


    move: IDLE ⇒ /eqP IDLE.

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

1 subgoal (ID 1409)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j (succn t)
  NCOMPL : ~~ completed_by sched j (succn t)
  IDLE : sched t = None
  ============================
  False

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


    move: (H_work_conserving j t ARR) ⇒ NIDLE.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1411)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j (succn t)
  NCOMPL : ~~ completed_by sched j (succn t)
  IDLE : sched t = None
  NIDLE : backlogged sched j t ->
          exists j_other : Job, scheduled_at sched j_other t
  ============================
  False

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


    feed NIDLE.

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

2 subgoals (ID 1412)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j (succn t)
  NCOMPL : ~~ completed_by sched j (succn t)
  IDLE : sched t = None
  NIDLE : backlogged sched j t ->
          exists j_other : Job, scheduled_at sched j_other t
  ============================
  backlogged sched j t

subgoal 2 (ID 1417) is:
 False

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


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

1 subgoal (ID 1412)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j (succn t)
  NCOMPL : ~~ completed_by sched j (succn t)
  IDLE : sched t = None
  NIDLE : backlogged sched j t ->
          exists j_other : Job, scheduled_at sched j_other t
  ============================
  backlogged sched j t

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


apply/andP; split; last first.

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

2 subgoals (ID 1444)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j (succn t)
  NCOMPL : ~~ completed_by sched j (succn t)
  IDLE : sched t = None
  NIDLE : backlogged sched j t ->
          exists j_other : Job, scheduled_at sched j_other t
  ============================
  ~~ scheduled_at sched j t

subgoal 2 (ID 1443) is:
 job_ready sched j t

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


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

1 subgoal (ID 1444)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j (succn t)
  NCOMPL : ~~ completed_by sched j (succn t)
  IDLE : sched t = None
  NIDLE : backlogged sched j t ->
          exists j_other : Job, scheduled_at sched j_other t
  ============================
  ~~ scheduled_at sched j t

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


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

2 subgoals (ID 1443)

subgoal 1 (ID 1443) is:
 job_ready sched j t
subgoal 2 (ID 1417) is:
 False

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


}

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

1 subgoal (ID 1443)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j (succn t)
  NCOMPL : ~~ completed_by sched j (succn t)
  IDLE : sched t = None
  NIDLE : backlogged sched j t ->
          exists j_other : Job, scheduled_at sched j_other t
  ============================
  job_ready sched j t

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


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

1 subgoal (ID 1443)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j (succn t)
  NCOMPL : ~~ completed_by sched j (succn t)
  IDLE : sched t = None
  NIDLE : backlogged sched j t ->
          exists j_other : Job, scheduled_at sched j_other t
  ============================
  job_ready sched j t

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


apply/andP; split; first by done.

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

1 subgoal (ID 1480)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j (succn t)
  NCOMPL : ~~ completed_by sched j (succn t)
  IDLE : sched t = None
  NIDLE : backlogged sched j t ->
          exists j_other : Job, scheduled_at sched j_other t
  ============================
  ~~ completed_by sched j t

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


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

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

1 subgoal (ID 1536)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j (succn t)
  IDLE : sched t = None
  NIDLE : backlogged sched j t ->
          exists j_other : Job, scheduled_at sched j_other t
  NC2 : completed_by sched j t
  ============================
  completed_by sched j (succn t)

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


          by apply completion_monotonic with t.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1417)

subgoal 1 (ID 1417) is:
 False

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


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

1 subgoal (ID 1417)

subgoal 1 (ID 1417) is:
 False

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


 
    }

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

1 subgoal (ID 1417)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j (succn t)
  NCOMPL : ~~ completed_by sched j (succn t)
  IDLE : sched t = None
  NIDLE : exists j_other : Job, scheduled_at sched j_other t
  ============================
  False

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


    move: NIDLE ⇒ [j' SCHED].

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

1 subgoal (ID 1555)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> Prop
  H_work_conserving : work_conserving arr_seq sched
  H_arrival_sequence_is_a_set : arrival_sequence_uniq arr_seq
  t : instant
  j : Job
  ARR : arrives_in arr_seq j
  HA : arrived_before j (succn t)
  NCOMPL : ~~ completed_by sched j (succn t)
  IDLE : sched t = None
  j' : Job
  SCHED : scheduled_at sched j' t
  ============================
  False

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


      by rewrite scheduled_at_def IDLE in SCHED.

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

No more subgoals.

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


  Qed.

Let the priority relation be reflexive.
Recall the notion of workload of all jobs released in a given interval [t1, t2)...
... and total service of jobs within some time interval [t1, t2).
Assume that for some positive Δ, the sum of requested workload at time (t + Δ) is bounded by delta (i.e., the supply). Note that this assumption bounds the total workload of jobs released in a time interval [t, t + Δ) regardless of their priorities.
  Variable Δ: duration.
  Hypothesis H_delta_positive: Δ > 0.
  Hypothesis H_workload_is_bounded: t, total_workload t (t + Δ) Δ.

Next we prove that since for any time instant t there is a point where the total workload is upper-bounded by the supply the processor encounters no carry-in instants at least every Δ time units.
  Section ProcessorIsNotTooBusy.

We start by proving that the processor has no carry-in at the beginning (i.e., has no carry-in at time instant 0).
    Lemma no_carry_in_at_the_beginning :
      no_carry_in 0.

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

1 subgoal (ID 1283)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  ============================
  no_carry_in 0

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


    Proof.
      intros s ARR AB; exfalso.

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

1 subgoal (ID 1288)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  s : Job
  ARR : arrives_in arr_seq s
  AB : arrived_before s 0
  ============================
  False

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


        by rewrite /arrived_before ltn0 in AB.

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

No more subgoals.

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


    Qed.

In this section, we prove that for any time instant t there exists another time instant t' ∈ (t, t + Δ] such that the processor has no carry-in at time t'.
Consider an arbitrary time instant t...
      Variable t: duration.

...such that the processor has no carry-in at time t.
      Hypothesis H_no_carry_in: no_carry_in t.

First, recall that the total service is bounded by the total workload. Therefore the total service of jobs in the interval [t, t + Δ) is bounded by Δ.
      Lemma total_service_is_bounded_by_Δ :
        total_service t (t + Δ) Δ.

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

1 subgoal (ID 1284)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  ============================
  total_service t (t + Δ) <= Δ

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


      Proof.
        unfold total_service.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1286)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  ============================
  service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) <= Δ

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


        rewrite -{3}[Δ]addn0 -{2}(subnn t) addnBA // [in X in _ X]addnC.

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

1 subgoal (ID 1333)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  ============================
  service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) <=
  t + Δ - t

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


        apply service_of_jobs_le_length_of_interval'.

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

1 subgoal (ID 1335)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  ============================
  uniq (arrivals_between 0 (t + Δ))

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


          by eapply arrivals_uniq; eauto 2.

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

No more subgoals.

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


      Qed.

Next we consider two cases: (1) The case when the total service is strictly less than Δ, (2) And when the total service is equal to Δ.
In the first case, we use the pigeonhole principle to conclude that there is an idle time instant; which in turn implies existence of a time instant with no carry-in.
      Lemma low_total_service_implies_existence_of_time_with_no_carry_in :
        total_service t (t + Δ) < Δ
         δ, δ < Δ no_carry_in (t.+1 + δ).

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

1 subgoal (ID 1287)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  ============================
  total_service t (t + Δ) < Δ ->
  exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


      Proof.
        unfold total_service; intros LT.

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

1 subgoal (ID 1290)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  LT : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) < Δ
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


        rewrite -{3}[Δ]addn0 -{2}(subnn t) addnBA // [Δ + t]addnC in LT.

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

1 subgoal (ID 1375)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  LT : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) <
       t + Δ - t
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


        eapply low_service_implies_existence_of_idle_time in LT; eauto.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1378)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  LT : exists t0 : nat, t <= t0 < t + Δ /\ is_idle sched t0
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


        move: LT ⇒ [t_idle [/andP [LEt GTe] IDLE]].

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

1 subgoal (ID 1543)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  t_idle : nat
  LEt : t <= t_idle
  GTe : t_idle < t + Δ
  IDLE : is_idle sched t_idle
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


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

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

2 subgoals (ID 1624)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  t_idle : nat
  GTe : t_idle < t + Δ
  IDLE : is_idle sched t_idle
  EQ : t = t_idle
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

subgoal 2 (ID 1625) is:
 exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


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

1 subgoal (ID 1624)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  t_idle : nat
  GTe : t_idle < t + Δ
  IDLE : is_idle sched t_idle
  EQ : t = t_idle
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


0; split; first done.

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

1 subgoal (ID 1630)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  t_idle : nat
  GTe : t_idle < t + Δ
  IDLE : is_idle sched t_idle
  EQ : t = t_idle
  ============================
  no_carry_in (succn t + 0)

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


          rewrite addn0; subst t_idle.

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

1 subgoal (ID 1643)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  IDLE : is_idle sched t
  GTe : t < t + Δ
  ============================
  no_carry_in (succn t)

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


          intros s ARR BEF.

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

1 subgoal (ID 1647)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  IDLE : is_idle sched t
  GTe : t < t + Δ
  s : Job
  ARR : arrives_in arr_seq s
  BEF : arrived_before s (succn t)
  ============================
  completed_by sched s (succn t)

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


          apply idle_instant_implies_no_carry_in_at_t_pl_1 in IDLE; try done.

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

1 subgoal (ID 1648)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  IDLE : no_carry_in (succn t)
  GTe : t < t + Δ
  s : Job
  ARR : arrives_in arr_seq s
  BEF : arrived_before s (succn t)
  ============================
  completed_by sched s (succn t)

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


            by apply IDLE.

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

1 subgoal (ID 1625)

subgoal 1 (ID 1625) is:
 exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


        }

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

1 subgoal (ID 1625)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  t_idle : nat
  GTe : t_idle < t + Δ
  IDLE : is_idle sched t_idle
  LT : t < t_idle
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


        have EX: γ, t_idle = t + γ.

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

2 subgoals (ID 1678)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  t_idle : nat
  GTe : t_idle < t + Δ
  IDLE : is_idle sched t_idle
  LT : t < t_idle
  ============================
  exists γ : nat, t_idle = t + γ

subgoal 2 (ID 1680) is:
 exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


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

1 subgoal (ID 1678)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  t_idle : nat
  GTe : t_idle < t + Δ
  IDLE : is_idle sched t_idle
  LT : t < t_idle
  ============================
  exists γ : nat, t_idle = t + γ

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


by (t_idle - t); rewrite subnKC // ltnW.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1680)

subgoal 1 (ID 1680) is:
 exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


}

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

1 subgoal (ID 1680)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  t_idle : nat
  GTe : t_idle < t + Δ
  IDLE : is_idle sched t_idle
  LT : t < t_idle
  EX : exists γ : nat, t_idle = t + γ
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


        move: EX ⇒ [γ EQ]; subst t_idle; rewrite ltn_add2l in GTe.

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

1 subgoal (ID 1783)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  γ : nat
  LT : t < t + γ
  IDLE : is_idle sched (t + γ)
  GTe : γ < Δ
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


        rewrite -{1}[t]addn0 ltn_add2l in LT.

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

1 subgoal (ID 1832)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  γ : nat
  IDLE : is_idle sched (t + γ)
  GTe : γ < Δ
  LT : 0 < γ
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


         (γ.-1); split.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1836)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  γ : nat
  IDLE : is_idle sched (t + γ)
  GTe : γ < Δ
  LT : 0 < γ
  ============================
  predn γ < Δ

subgoal 2 (ID 1837) is:
 no_carry_in (succn t + predn γ)

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


        - apply leq_trans with γ.
(* ----------------------------------[ coqtop ]---------------------------------

3 subgoals (ID 1838)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  γ : nat
  IDLE : is_idle sched (t + γ)
  GTe : γ < Δ
  LT : 0 < γ
  ============================
  predn γ < γ

subgoal 2 (ID 1839) is:
 γ <= Δ
subgoal 3 (ID 1837) is:
 no_carry_in (succn t + predn γ)

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


by rewrite prednK.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1839)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  γ : nat
  IDLE : is_idle sched (t + γ)
  GTe : γ < Δ
  LT : 0 < γ
  ============================
  γ <= Δ

subgoal 2 (ID 1837) is:
 no_carry_in (succn t + predn γ)

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


by rewrite ltnW.

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

1 subgoal (ID 1837)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  γ : nat
  IDLE : is_idle sched (t + γ)
  GTe : γ < Δ
  LT : 0 < γ
  ============================
  no_carry_in (succn t + predn γ)

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


        - rewrite -subn1 -addn1 -addnA subnKC //.

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

1 subgoal (ID 1868)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  γ : nat
  IDLE : is_idle sched (t + γ)
  GTe : γ < Δ
  LT : 0 < γ
  ============================
  no_carry_in (t + γ)

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


          intros s ARR BEF.

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

1 subgoal (ID 1898)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  γ : nat
  IDLE : is_idle sched (t + γ)
  GTe : γ < Δ
  LT : 0 < γ
  s : Job
  ARR : arrives_in arr_seq s
  BEF : arrived_before s (t + γ)
  ============================
  completed_by sched s (t + γ)

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


            by apply idle_instant_implies_no_carry_in_at_t.

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

No more subgoals.

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


      Qed.

In the second case, the total service within the time interval [t, t + Δ) is equal to Δ. On the other hand, we know that the total workload is lower-bounded by the total service and upper-bounded by Δ. Therefore, the total workload is equal to total service this implies completion of all jobs by time [t + Δ] and hence no carry-in at time [t + Δ].
      Lemma completion_of_all_jobs_implies_no_carry_in :
        total_service t (t + Δ) = Δ
        no_carry_in (t + Δ).

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

1 subgoal (ID 1290)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  ============================
  total_service t (t + Δ) = Δ -> no_carry_in (t + Δ)

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


      Proof.
        unfold total_service; intros EQserv.

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

1 subgoal (ID 1293)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  ============================
  no_carry_in (t + Δ)

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


        move: (H_workload_is_bounded t); moveWORK.

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

1 subgoal (ID 1295)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  ============================
  no_carry_in (t + Δ)

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


        have EQ: total_workload 0 (t + Δ) = service_of_jobs sched predT (arrivals_between 0 (t + Δ)) 0 (t + Δ).

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

2 subgoals (ID 1302)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  ============================
  total_workload 0 (t + Δ) =
  service_of_jobs sched predT (arrivals_between 0 (t + Δ)) 0 (t + Δ)

subgoal 2 (ID 1304) is:
 no_carry_in (t + Δ)

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


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

1 subgoal (ID 1302)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  ============================
  total_workload 0 (t + Δ) =
  service_of_jobs sched predT (arrivals_between 0 (t + Δ)) 0 (t + Δ)

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


intros.
          have COMPL := all_jobs_have_completed_impl_workload_eq_service
                          arr_seq H_arrival_times_are_consistent sched
                          H_jobs_must_arrive_to_execute
                          H_completed_jobs_dont_execute
                          predT 0 t t.

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

1 subgoal (ID 1313)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  COMPL : (forall j : Job,
           j \in arrival_sequence.arrivals_between arr_seq 0 t ->
           predT j -> completed_by sched j t) ->
          workload_of_jobs predT
            (arrival_sequence.arrivals_between arr_seq 0 t) =
          service_of_jobs sched predT
            (arrival_sequence.arrivals_between arr_seq 0 t) 0 t
  ============================
  total_workload 0 (t + Δ) =
  service_of_jobs sched predT (arrivals_between 0 (t + Δ)) 0 (t + Δ)

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


          feed COMPL; try done.

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

2 subgoals (ID 1314)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  COMPL : (forall j : Job,
           j \in arrival_sequence.arrivals_between arr_seq 0 t ->
           predT j -> completed_by sched j t) ->
          workload_of_jobs predT
            (arrival_sequence.arrivals_between arr_seq 0 t) =
          service_of_jobs sched predT
            (arrival_sequence.arrivals_between arr_seq 0 t) 0 t
  ============================
  forall j : Job,
  j \in arrival_sequence.arrivals_between arr_seq 0 t ->
  predT j -> completed_by sched j t

subgoal 2 (ID 1319) is:
 total_workload 0 (t + Δ) =
 service_of_jobs sched predT (arrivals_between 0 (t + Δ)) 0 (t + Δ)

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


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

1 subgoal (ID 1314)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  COMPL : (forall j : Job,
           j \in arrival_sequence.arrivals_between arr_seq 0 t ->
           predT j -> completed_by sched j t) ->
          workload_of_jobs predT
            (arrival_sequence.arrivals_between arr_seq 0 t) =
          service_of_jobs sched predT
            (arrival_sequence.arrivals_between arr_seq 0 t) 0 t
  ============================
  forall j : Job,
  j \in arrival_sequence.arrivals_between arr_seq 0 t ->
  predT j -> completed_by sched j t

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


intros j A B; apply H_no_carry_in.

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

2 subgoals (ID 1373)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  COMPL : (forall j : Job,
           j \in arrival_sequence.arrivals_between arr_seq 0 t ->
           predT j -> completed_by sched j t) ->
          workload_of_jobs predT
            (arrival_sequence.arrivals_between arr_seq 0 t) =
          service_of_jobs sched predT
            (arrival_sequence.arrivals_between arr_seq 0 t) 0 t
  j : Job
  A : j \in arrival_sequence.arrivals_between arr_seq 0 t
  B : predT j
  ============================
  arrives_in arr_seq j

subgoal 2 (ID 1374) is:
 arrived_before j t

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


            - eapply in_arrivals_implies_arrived; eauto 2.

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

1 subgoal (ID 1374)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  COMPL : (forall j : Job,
           j \in arrival_sequence.arrivals_between arr_seq 0 t ->
           predT j -> completed_by sched j t) ->
          workload_of_jobs predT
            (arrival_sequence.arrivals_between arr_seq 0 t) =
          service_of_jobs sched predT
            (arrival_sequence.arrivals_between arr_seq 0 t) 0 t
  j : Job
  A : j \in arrival_sequence.arrivals_between arr_seq 0 t
  B : predT j
  ============================
  arrived_before j t

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


            - eapply in_arrivals_implies_arrived_between in A; eauto 2.

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

2 subgoals (ID 1319)

subgoal 1 (ID 1319) is:
 total_workload 0 (t + Δ) =
 service_of_jobs sched predT (arrivals_between 0 (t + Δ)) 0 (t + Δ)
subgoal 2 (ID 1304) is:
 no_carry_in (t + Δ)

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


          }

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

1 subgoal (ID 1319)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  COMPL : workload_of_jobs predT
            (arrival_sequence.arrivals_between arr_seq 0 t) =
          service_of_jobs sched predT
            (arrival_sequence.arrivals_between arr_seq 0 t) 0 t
  ============================
  total_workload 0 (t + Δ) =
  service_of_jobs sched predT (arrivals_between 0 (t + Δ)) 0 (t + Δ)

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


          apply/eqP; rewrite eqn_leq; apply/andP; split;
            last by apply service_of_jobs_le_workload;
            auto using ideal_proc_model_provides_unit_service.

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

1 subgoal (ID 1472)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  COMPL : workload_of_jobs predT
            (arrival_sequence.arrivals_between arr_seq 0 t) =
          service_of_jobs sched predT
            (arrival_sequence.arrivals_between arr_seq 0 t) 0 t
  ============================
  total_workload 0 (t + Δ) <=
  service_of_jobs sched predT (arrivals_between 0 (t + Δ)) 0 (t + Δ)

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


          rewrite /total_workload (workload_of_jobs_cat arr_seq t); last first.

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

2 subgoals (ID 1493)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  COMPL : workload_of_jobs predT
            (arrival_sequence.arrivals_between arr_seq 0 t) =
          service_of_jobs sched predT
            (arrival_sequence.arrivals_between arr_seq 0 t) 0 t
  ============================
  0 <= t <= t + Δ

subgoal 2 (ID 1492) is:
 workload_of_jobs predT (arrival_sequence.arrivals_between arr_seq 0 t) +
 workload_of_jobs predT (arrival_sequence.arrivals_between arr_seq t (t + Δ)) <=
 service_of_jobs sched predT (arrivals_between 0 (t + Δ)) 0 (t + Δ)

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


          apply/andP; split; [by done | by rewrite leq_addr].

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

1 subgoal (ID 1492)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  COMPL : workload_of_jobs predT
            (arrival_sequence.arrivals_between arr_seq 0 t) =
          service_of_jobs sched predT
            (arrival_sequence.arrivals_between arr_seq 0 t) 0 t
  ============================
  workload_of_jobs predT (arrival_sequence.arrivals_between arr_seq 0 t) +
  workload_of_jobs predT
    (arrival_sequence.arrivals_between arr_seq t (t + Δ)) <=
  service_of_jobs sched predT (arrivals_between 0 (t + Δ)) 0 (t + Δ)

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


          rewrite (service_of_jobs_cat_scheduling_interval _ _ _ _ _ _ _ t); try done; first last.

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

2 subgoals (ID 1566)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  COMPL : workload_of_jobs predT
            (arrival_sequence.arrivals_between arr_seq 0 t) =
          service_of_jobs sched predT
            (arrival_sequence.arrivals_between arr_seq 0 t) 0 t
  ============================
  0 <= t <= t + Δ

subgoal 2 (ID 1563) is:
 workload_of_jobs predT (arrival_sequence.arrivals_between arr_seq 0 t) +
 workload_of_jobs predT (arrival_sequence.arrivals_between arr_seq t (t + Δ)) <=
 service_of_jobs sched predT (arrival_sequence.arrivals_between arr_seq 0 t)
   0 t +
 service_of_jobs sched predT (arrival_sequence.arrivals_between arr_seq 0 t)
   t (t + Δ) +
 service_of_jobs sched predT
   (arrival_sequence.arrivals_between arr_seq t (t + Δ)) t 
   (t + Δ)

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


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

1 subgoal (ID 1566)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  COMPL : workload_of_jobs predT
            (arrival_sequence.arrivals_between arr_seq 0 t) =
          service_of_jobs sched predT
            (arrival_sequence.arrivals_between arr_seq 0 t) 0 t
  ============================
  0 <= t <= t + Δ

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


by apply/andP; split; [by done | by rewrite leq_addr].
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1563)

subgoal 1 (ID 1563) is:
 workload_of_jobs predT (arrival_sequence.arrivals_between arr_seq 0 t) +
 workload_of_jobs predT (arrival_sequence.arrivals_between arr_seq t (t + Δ)) <=
 service_of_jobs sched predT (arrival_sequence.arrivals_between arr_seq 0 t)
   0 t +
 service_of_jobs sched predT (arrival_sequence.arrivals_between arr_seq 0 t)
   t (t + Δ) +
 service_of_jobs sched predT
   (arrival_sequence.arrivals_between arr_seq t (t + Δ)) t 
   (t + Δ)
subgoal 2 (ID 1304) is:
 no_carry_in (t + Δ)

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


}

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

1 subgoal (ID 1563)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  COMPL : workload_of_jobs predT
            (arrival_sequence.arrivals_between arr_seq 0 t) =
          service_of_jobs sched predT
            (arrival_sequence.arrivals_between arr_seq 0 t) 0 t
  ============================
  workload_of_jobs predT (arrival_sequence.arrivals_between arr_seq 0 t) +
  workload_of_jobs predT
    (arrival_sequence.arrivals_between arr_seq t (t + Δ)) <=
  service_of_jobs sched predT (arrival_sequence.arrivals_between arr_seq 0 t)
    0 t +
  service_of_jobs sched predT (arrival_sequence.arrivals_between arr_seq 0 t)
    t (t + Δ) +
  service_of_jobs sched predT
    (arrival_sequence.arrivals_between arr_seq t (t + Δ)) t 
    (t + Δ)

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


          rewrite COMPL -addnA leq_add2l.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1655)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  COMPL : workload_of_jobs predT
            (arrival_sequence.arrivals_between arr_seq 0 t) =
          service_of_jobs sched predT
            (arrival_sequence.arrivals_between arr_seq 0 t) 0 t
  ============================
  workload_of_jobs predT
    (arrival_sequence.arrivals_between arr_seq t (t + Δ)) <=
  service_of_jobs sched predT (arrival_sequence.arrivals_between arr_seq 0 t)
    t (t + Δ) +
  service_of_jobs sched predT
    (arrival_sequence.arrivals_between arr_seq t (t + Δ)) t 
    (t + Δ)

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


          rewrite -service_of_jobs_cat_arrival_interval; last first.

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

2 subgoals (ID 1671)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  COMPL : workload_of_jobs predT
            (arrival_sequence.arrivals_between arr_seq 0 t) =
          service_of_jobs sched predT
            (arrival_sequence.arrivals_between arr_seq 0 t) 0 t
  ============================
  0 <= t <= t + Δ

subgoal 2 (ID 1670) is:
 workload_of_jobs predT (arrival_sequence.arrivals_between arr_seq t (t + Δ)) <=
 service_of_jobs sched predT
   (arrival_sequence.arrivals_between arr_seq 0 (t + Δ)) t 
   (t + Δ)

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


          apply/andP; split; [by done| by rewrite leq_addr].

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

1 subgoal (ID 1670)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  COMPL : workload_of_jobs predT
            (arrival_sequence.arrivals_between arr_seq 0 t) =
          service_of_jobs sched predT
            (arrival_sequence.arrivals_between arr_seq 0 t) 0 t
  ============================
  workload_of_jobs predT
    (arrival_sequence.arrivals_between arr_seq t (t + Δ)) <=
  service_of_jobs sched predT
    (arrival_sequence.arrivals_between arr_seq 0 (t + Δ)) t 
    (t + Δ)

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


          rewrite EQserv.

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

1 subgoal (ID 1704)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  COMPL : workload_of_jobs predT
            (arrival_sequence.arrivals_between arr_seq 0 t) =
          service_of_jobs sched predT
            (arrival_sequence.arrivals_between arr_seq 0 t) 0 t
  ============================
  workload_of_jobs predT
    (arrival_sequence.arrivals_between arr_seq t (t + Δ)) <= Δ

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


            by apply H_workload_is_bounded.

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

1 subgoal (ID 1304)

subgoal 1 (ID 1304) is:
 no_carry_in (t + Δ)

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


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

1 subgoal (ID 1304)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  EQ : total_workload 0 (t + Δ) =
       service_of_jobs sched predT (arrivals_between 0 (t + Δ)) 0 (t + Δ)
  ============================
  no_carry_in (t + Δ)

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


 
        intros s ARR BEF.

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

1 subgoal (ID 1708)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  EQ : total_workload 0 (t + Δ) =
       service_of_jobs sched predT (arrivals_between 0 (t + Δ)) 0 (t + Δ)
  s : Job
  ARR : arrives_in arr_seq s
  BEF : arrived_before s (t + Δ)
  ============================
  completed_by sched s (t + Δ)

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


        eapply workload_eq_service_impl_all_jobs_have_completed; eauto 2; try done.

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

1 subgoal (ID 1720)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : duration
  H_no_carry_in : no_carry_in t
  EQserv : service_of_jobs sched predT (arrivals_between 0 (t + Δ)) t (t + Δ) =
           Δ
  WORK : total_workload t (t + Δ) <= Δ
  EQ : total_workload 0 (t + Δ) =
       service_of_jobs sched predT (arrivals_between 0 (t + Δ)) 0 (t + Δ)
  s : Job
  ARR : arrives_in arr_seq s
  BEF : arrived_before s (t + Δ)
  ============================
  s \in arrival_sequence.arrivals_between arr_seq 0 (t + Δ)

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


          by eapply arrived_between_implies_in_arrivals; eauto 2.

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

No more subgoals.

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


      Qed.

    End ProcessorIsNotTooBusyInduction.

Finally, we show that any interval of length Δ contains a time instant with no carry-in.
    Lemma processor_is_not_too_busy :
       t, δ, δ < Δ no_carry_in (t + δ).

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

1 subgoal (ID 1287)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  ============================
  forall t : nat, exists δ : nat, δ < Δ /\ no_carry_in (t + δ)

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


    Proof.
      induction t.

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

2 subgoals (ID 1292)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (0 + δ)

subgoal 2 (ID 1295) is:
 exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


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

1 subgoal (ID 1292)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (0 + δ)

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


by 0; split; [ | rewrite addn0; apply no_carry_in_at_the_beginning].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1295)

subgoal 1 (ID 1295) is:
 exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


}

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

1 subgoal (ID 1295)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : nat
  IHt : exists δ : nat, δ < Δ /\ no_carry_in (t + δ)
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


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

1 subgoal (ID 1295)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : nat
  IHt : exists δ : nat, δ < Δ /\ no_carry_in (t + δ)
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


move: IHt ⇒ [δ [LE FQT]].

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

1 subgoal (ID 1328)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t, δ : nat
  LE : δ < Δ
  FQT : no_carry_in (t + δ)
  ============================
  exists δ0 : nat, δ0 < Δ /\ no_carry_in (succn t + δ0)

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


        move: (posnP δ) ⇒ [Z|POS]; last first.

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

2 subgoals (ID 1344)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t, δ : nat
  LE : δ < Δ
  FQT : no_carry_in (t + δ)
  POS : 0 < δ
  ============================
  exists δ0 : nat, δ0 < Δ /\ no_carry_in (succn t + δ0)

subgoal 2 (ID 1343) is:
 exists δ0 : nat, δ0 < Δ /\ no_carry_in (succn t + δ0)

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


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

1 subgoal (ID 1344)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t, δ : nat
  LE : δ < Δ
  FQT : no_carry_in (t + δ)
  POS : 0 < δ
  ============================
  exists δ0 : nat, δ0 < Δ /\ no_carry_in (succn t + δ0)

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


(δ.-1); split.

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

2 subgoals (ID 1348)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t, δ : nat
  LE : δ < Δ
  FQT : no_carry_in (t + δ)
  POS : 0 < δ
  ============================
  predn δ < Δ

subgoal 2 (ID 1349) is:
 no_carry_in (succn t + predn δ)

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


          - by apply leq_trans with δ; [rewrite prednK | apply ltnW].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1349)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t, δ : nat
  LE : δ < Δ
  FQT : no_carry_in (t + δ)
  POS : 0 < δ
  ============================
  no_carry_in (succn t + predn δ)

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


          - by rewrite -subn1 -addn1 -addnA subnKC //.

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

1 subgoal (ID 1343)

subgoal 1 (ID 1343) is:
 exists δ0 : nat, δ0 < Δ /\ no_carry_in (succn t + δ0)

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


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

1 subgoal (ID 1343)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t, δ : nat
  LE : δ < Δ
  FQT : no_carry_in (t + δ)
  Z : δ = 0
  ============================
  exists δ0 : nat, δ0 < Δ /\ no_carry_in (succn t + δ0)

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


subst δ; rewrite addn0 in FQT; clear LE.

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

1 subgoal (ID 1426)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : nat
  FQT : no_carry_in t
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


        move: (total_service_is_bounded_by_Δ t); rewrite leq_eqVlt; move ⇒ /orP [/eqP EQ | LT].

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

2 subgoals (ID 1506)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : nat
  FQT : no_carry_in t
  EQ : total_service t (t + Δ) = Δ
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

subgoal 2 (ID 1507) is:
 exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


        - (Δ.-1); split.

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

3 subgoals (ID 1511)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : nat
  FQT : no_carry_in t
  EQ : total_service t (t + Δ) = Δ
  ============================
  predn Δ < Δ

subgoal 2 (ID 1512) is:
 no_carry_in (succn t + predn Δ)
subgoal 3 (ID 1507) is:
 exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


          + by rewrite prednK.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1512)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : nat
  FQT : no_carry_in t
  EQ : total_service t (t + Δ) = Δ
  ============================
  no_carry_in (succn t + predn Δ)

subgoal 2 (ID 1507) is:
 exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


          + by rewrite addSn -subn1 -addn1 -addnA subnK; first apply completion_of_all_jobs_implies_no_carry_in.

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

1 subgoal (ID 1507)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  t : nat
  FQT : no_carry_in t
  LT : total_service t (t + Δ) < Δ
  ============================
  exists δ : nat, δ < Δ /\ no_carry_in (succn t + δ)

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


        - by apply low_total_service_implies_existence_of_time_with_no_carry_in.
(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

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


      }

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

No more subgoals.

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


    Qed.

  End ProcessorIsNotTooBusy.

Consider an arbitrary job j with positive cost.
  Variable j : Job.
  Hypothesis H_from_arrival_sequence : arrives_in arr_seq j.
  Hypothesis H_job_cost_positive : job_cost_positive j.

We show that there must exist a busy interval [t1, t2) that contains the arrival time of j.
  Corollary exists_busy_interval_from_total_workload_bound :
     t1 t2,
      t1 job_arrival j < t2
      t2 t1 + Δ
      busy_interval arr_seq sched j t1 t2.

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

1 subgoal (ID 1300)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  H_from_arrival_sequence : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  ============================
  exists t1 t2 : nat,
    t1 <= job_arrival j < t2 /\
    t2 <= t1 + Δ /\ busy_interval arr_seq sched j t1 t2

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


  Proof.
    rename H_from_arrival_sequence into ARR, H_job_cost_positive into POS.

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

1 subgoal (ID 1301)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  ============================
  exists t1 t2 : nat,
    t1 <= job_arrival j < t2 /\
    t2 <= t1 + Δ /\ busy_interval arr_seq sched j t1 t2

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


    edestruct (exists_busy_interval_prefix
                 arr_seq H_arrival_times_are_consistent sched j ARR H_priority_is_reflexive (job_arrival j))
      as [t1 [PREFIX GE]].

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

2 subgoals (ID 1309)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  ============================
  pending sched j (job_arrival j)

subgoal 2 (ID 1319) is:
 exists t0 t2 : nat,
   t0 <= job_arrival j < t2 /\
   t2 <= t0 + Δ /\ busy_interval arr_seq sched j t0 t2

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


    apply job_pending_at_arrival; auto.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1319)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE : t1 <= job_arrival j <= job_arrival j
  ============================
  exists t0 t2 : nat,
    t0 <= job_arrival j < t2 /\
    t2 <= t0 + Δ /\ busy_interval arr_seq sched j t0 t2

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


    move: GE ⇒ /andP [GE1 _].

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

1 subgoal (ID 1368)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  ============================
  exists t0 t2 : nat,
    t0 <= job_arrival j < t2 /\
    t2 <= t0 + Δ /\ busy_interval arr_seq sched j t0 t2

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


     t1; move: (processor_is_not_too_busy t1.+1) ⇒ [δ [LE QT]].

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

1 subgoal (ID 1392)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : no_carry_in (succn t1 + δ)
  ============================
  exists t2 : nat,
    t1 <= job_arrival j < t2 /\
    t2 <= t1 + Δ /\ busy_interval arr_seq sched j t1 t2

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


    apply no_carry_in_implies_quiet_time with (j := j) in QT.

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

1 subgoal (ID 1393)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  ============================
  exists t2 : nat,
    t1 <= job_arrival j < t2 /\
    t2 <= t1 + Δ /\ busy_interval arr_seq sched j t1 t2

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


    have EX: t2, ((t1 < t2 t1.+1 + δ) && quiet_time_dec arr_seq sched j t2).

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

2 subgoals (ID 1402)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  ============================
  exists t2 : nat,
    (t1 < t2 <= succn t1 + δ) && quiet_time_dec arr_seq sched j t2

subgoal 2 (ID 1404) is:
 exists t2 : nat,
   t1 <= job_arrival j < t2 /\
   t2 <= t1 + Δ /\ busy_interval arr_seq sched j t1 t2

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


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

1 subgoal (ID 1402)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  ============================
  exists t2 : nat,
    (t1 < t2 <= succn t1 + δ) && quiet_time_dec arr_seq sched j t2

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


(t1.+1 + δ); apply/andP; split.

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

2 subgoals (ID 1432)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  ============================
  t1 < succn t1 + δ <= succn t1 + δ

subgoal 2 (ID 1433) is:
 quiet_time_dec arr_seq sched j (succn t1 + δ)

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


      - by apply/andP; split; first rewrite addSn ltnS leq_addr.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1433)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  ============================
  quiet_time_dec arr_seq sched j (succn t1 + δ)

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


      - by apply/quiet_time_P.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1404)

subgoal 1 (ID 1404) is:
 exists t2 : nat,
   t1 <= job_arrival j < t2 /\
   t2 <= t1 + Δ /\ busy_interval arr_seq sched j t1 t2

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


}

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

1 subgoal (ID 1404)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  EX : exists t2 : nat,
         (t1 < t2 <= succn t1 + δ) && quiet_time_dec arr_seq sched j t2
  ============================
  exists t2 : nat,
    t1 <= job_arrival j < t2 /\
    t2 <= t1 + Δ /\ busy_interval arr_seq sched j t1 t2

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


    move: (ex_minnP EX) ⇒ [t2 /andP [/andP [GTt2 LEt2] QUIET] MIN]; clear EX.

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

1 subgoal (ID 1619)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  ============================
  exists t0 : nat,
    t1 <= job_arrival j < t0 /\
    t0 <= t1 + Δ /\ busy_interval arr_seq sched j t1 t0

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


    have NEQ: t1 job_arrival j < t2.

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

2 subgoals (ID 1624)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  ============================
  t1 <= job_arrival j < t2

subgoal 2 (ID 1626) is:
 exists t0 : nat,
   t1 <= job_arrival j < t0 /\
   t0 <= t1 + Δ /\ busy_interval arr_seq sched j t1 t0

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


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

1 subgoal (ID 1624)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  ============================
  t1 <= job_arrival j < t2

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


apply/andP; split; first by done.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1653)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  ============================
  job_arrival j < t2

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


      rewrite ltnNge; apply/negP; intros CONTR.

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

1 subgoal (ID 1679)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  CONTR : t2 <= job_arrival j
  ============================
  False

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


      move: (PREFIX) ⇒ [_ [_ [NQT _]]].

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

1 subgoal (ID 1712)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  CONTR : t2 <= job_arrival j
  NQT : forall t : nat,
        t1 < t < succn (job_arrival j) ->
        ~ busy_interval.quiet_time arr_seq sched j t
  ============================
  False

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


      move: (NQT t2); clear NQT; moveNQT.

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

1 subgoal (ID 1715)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  CONTR : t2 <= job_arrival j
  NQT : t1 < t2 < succn (job_arrival j) ->
        ~ busy_interval.quiet_time arr_seq sched j t2
  ============================
  False

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


      feed NQT; first by (apply/andP; split; [|rewrite ltnS]).
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1721)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  CONTR : t2 <= job_arrival j
  NQT : ~ busy_interval.quiet_time arr_seq sched j t2
  ============================
  False

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


        by apply: NQT; apply/quiet_time_P.

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

1 subgoal (ID 1626)

subgoal 1 (ID 1626) is:
 exists t0 : nat,
   t1 <= job_arrival j < t0 /\
   t0 <= t1 + Δ /\ busy_interval arr_seq sched j t1 t0

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


    }

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

1 subgoal (ID 1626)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  NEQ : t1 <= job_arrival j < t2
  ============================
  exists t0 : nat,
    t1 <= job_arrival j < t0 /\
    t0 <= t1 + Δ /\ busy_interval arr_seq sched j t1 t0

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


     t2; split; last split; first by done.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1871)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  NEQ : t1 <= job_arrival j < t2
  ============================
  t2 <= t1 + Δ

subgoal 2 (ID 1872) is:
 busy_interval arr_seq sched j t1 t2

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


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

1 subgoal (ID 1871)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  NEQ : t1 <= job_arrival j < t2
  ============================
  t2 <= t1 + Δ

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


apply leq_trans with (t1.+1 + δ); [by done | by rewrite addSn ltn_add2l].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1872)

subgoal 1 (ID 1872) is:
 busy_interval arr_seq sched j t1 t2

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


}

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

1 subgoal (ID 1872)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  NEQ : t1 <= job_arrival j < t2
  ============================
  busy_interval arr_seq sched j t1 t2

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


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

1 subgoal (ID 1872)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  PREFIX : busy_interval_prefix arr_seq sched j t1 (succn (job_arrival j))
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  NEQ : t1 <= job_arrival j < t2
  ============================
  busy_interval arr_seq sched j t1 t2

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


move: PREFIX ⇒ [_ [QTt1 [NQT _]]]; repeat split; try done.

(* ----------------------------------[ 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
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  NEQ : t1 <= job_arrival j < t2
  QTt1 : busy_interval.quiet_time arr_seq sched j t1
  NQT : forall t : nat,
        t1 < t < succn (job_arrival j) ->
        ~ busy_interval.quiet_time arr_seq sched j t
  ============================
  forall t : nat, t1 < t < t2 -> ~ busy_interval.quiet_time arr_seq sched j t

subgoal 2 (ID 1920) is:
 busy_interval.quiet_time arr_seq sched j t2

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


      - movet /andP [GEt LTt] QTt.

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

2 subgoals (ID 2042)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  NEQ : t1 <= job_arrival j < t2
  QTt1 : busy_interval.quiet_time arr_seq sched j t1
  NQT : forall t : nat,
        t1 < t < succn (job_arrival j) ->
        ~ busy_interval.quiet_time arr_seq sched j t
  t : nat
  GEt : t1 < t
  LTt : t < t2
  QTt : busy_interval.quiet_time arr_seq sched j t
  ============================
  False

subgoal 2 (ID 1920) is:
 busy_interval.quiet_time arr_seq sched j t2

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


        feed (MIN t).

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

3 subgoals (ID 2043)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  NEQ : t1 <= job_arrival j < t2
  QTt1 : busy_interval.quiet_time arr_seq sched j t1
  NQT : forall t : nat,
        t1 < t < succn (job_arrival j) ->
        ~ busy_interval.quiet_time arr_seq sched j t
  t : nat
  GEt : t1 < t
  LTt : t < t2
  QTt : busy_interval.quiet_time arr_seq sched j t
  ============================
  (t1 < t <= succn t1 + δ) && quiet_time_dec arr_seq sched j t

subgoal 2 (ID 2048) is:
 False
subgoal 3 (ID 1920) is:
 busy_interval.quiet_time arr_seq sched j t2

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


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

1 subgoal (ID 2043)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  NEQ : t1 <= job_arrival j < t2
  QTt1 : busy_interval.quiet_time arr_seq sched j t1
  NQT : forall t : nat,
        t1 < t < succn (job_arrival j) ->
        ~ busy_interval.quiet_time arr_seq sched j t
  t : nat
  GEt : t1 < t
  LTt : t < t2
  QTt : busy_interval.quiet_time arr_seq sched j t
  ============================
  (t1 < t <= succn t1 + δ) && quiet_time_dec arr_seq sched j t

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


apply/andP; split.

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

2 subgoals (ID 2074)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  NEQ : t1 <= job_arrival j < t2
  QTt1 : busy_interval.quiet_time arr_seq sched j t1
  NQT : forall t : nat,
        t1 < t < succn (job_arrival j) ->
        ~ busy_interval.quiet_time arr_seq sched j t
  t : nat
  GEt : t1 < t
  LTt : t < t2
  QTt : busy_interval.quiet_time arr_seq sched j t
  ============================
  t1 < t <= succn t1 + δ

subgoal 2 (ID 2075) is:
 quiet_time_dec arr_seq sched j t

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


          + by apply/andP; split; last (apply leq_trans with t2; [apply ltnW | ]).

(* ----------------------------------[ 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
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  NEQ : t1 <= job_arrival j < t2
  QTt1 : busy_interval.quiet_time arr_seq sched j t1
  NQT : forall t : nat,
        t1 < t < succn (job_arrival j) ->
        ~ busy_interval.quiet_time arr_seq sched j t
  t : nat
  GEt : t1 < t
  LTt : t < t2
  QTt : busy_interval.quiet_time arr_seq sched j t
  ============================
  quiet_time_dec arr_seq sched j t

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


          + by apply/quiet_time_P.

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

2 subgoals (ID 2048)

subgoal 1 (ID 2048) is:
 False
subgoal 2 (ID 1920) is:
 busy_interval.quiet_time arr_seq sched j t2

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


        }

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

2 subgoals (ID 2048)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  job_pending_at := pending sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  t : nat
  MIN : t2 <= t
  NEQ : t1 <= job_arrival j < t2
  QTt1 : busy_interval.quiet_time arr_seq sched j t1
  NQT : forall t : nat,
        t1 < t < succn (job_arrival j) ->
        ~ busy_interval.quiet_time arr_seq sched j t
  GEt : t1 < t
  LTt : t < t2
  QTt : busy_interval.quiet_time arr_seq sched j t
  ============================
  False

subgoal 2 (ID 1920) is:
 busy_interval.quiet_time arr_seq sched j t2

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


          by move: LTt; rewrite ltnNge; move ⇒ /negP LTt; apply: LTt.

(* ----------------------------------[ 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
  no_carry_in := carry_in.no_carry_in arr_seq sched : instant -> Prop
  quiet_time := busy_interval.quiet_time arr_seq sched
   : Job -> instant -> 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
  total_workload := fun t1 t2 : instant =>
                    workload_of_jobs predT (arrivals_between t1 t2)
   : instant -> instant -> nat
  total_service := fun t1 t2 : instant =>
                   service_of_jobs sched predT (arrivals_between 0 t2) t1 t2
   : instant -> instant -> nat
  Δ : duration
  H_delta_positive : 0 < Δ
  H_workload_is_bounded : forall t : instant, total_workload t (t + Δ) <= Δ
  j : Job
  ARR : arrives_in arr_seq j
  POS : job_cost_positive j
  t1 : instant
  GE1 : t1 <= job_arrival j
  δ : nat
  LE : δ < Δ
  QT : quiet_time j (succn t1 + δ)
  t2 : nat
  GTt2 : t1 < t2
  LEt2 : t2 <= succn t1 + δ
  QUIET : quiet_time_dec arr_seq sched j t2
  MIN : forall n : nat,
        (t1 < n <= succn t1 + δ) && quiet_time_dec arr_seq sched j n ->
        t2 <= n
  NEQ : t1 <= job_arrival j < t2
  QTt1 : busy_interval.quiet_time arr_seq sched j t1
  NQT : forall t : nat,
        t1 < t < succn (job_arrival j) ->
        ~ busy_interval.quiet_time arr_seq sched j t
  ============================
  busy_interval.quiet_time arr_seq sched j t2

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


      - by apply/quiet_time_P.
(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

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


    }

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

No more subgoals.

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


  Qed.

End ExistsNoCarryIn.