Library prosa.analysis.facts.model.service_of_jobs


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

Welcome to Coq 8.13.0 (January 2021)

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


Require Export prosa.model.aggregate.workload.
Require Export prosa.model.aggregate.service_of_jobs.
Require Export prosa.analysis.facts.behavior.completion.
Require Export prosa.analysis.facts.model.ideal_schedule.
Require Import prosa.model.processor.ideal.

From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq fintype bigop.

Lemmas about Service Received by Sets of Jobs

In this file, we establish basic facts about the service received by _sets_ of jobs.
To begin with, we provide some basic properties of service of a set of jobs in case of a generic scheduling model.
Consider any type of tasks ...
  Context {Task : TaskType}.

... and any type of jobs associated with these tasks.
  Context {Job : JobType}.
  Context `{JobTask Job Task}.
  Context `{JobArrival Job}.
  Context `{JobCost Job}.

Consider any kind of processor state model, ...
  Context {PState : Type}.
  Context `{ProcessorState Job PState}.

... any job arrival sequence with consistent arrivals, ....
... and any schedule of this arrival sequence ...
  Variable sched : schedule PState.

... where jobs do not execute before their arrival or after completion.
Let P be any predicate over jobs.
  Variable P : Job bool.

Let's define some local names for clarity.
In this section, we prove that the service received by any set of jobs is upper-bounded by the corresponding workload.
Let jobs denote any (finite) set of jobs.
    Variable jobs : seq Job.

Assume that the processor model is a unit service model. I.e., no job ever receives more than one unit of service at any time.
Then, we prove that the service received by those jobs is no larger than their workload.
    Lemma service_of_jobs_le_workload:
       t1 t2,
        service_of_jobs sched P jobs t1 t2 workload_of_jobs P jobs.

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

1 subgoal (ID 38)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  jobs : seq Job
  H_unit_service : unit_service_proc_model PState
  ============================
  forall t1 t2 : instant,
  service_of_jobs sched P jobs t1 t2 <= workload_of_jobs P jobs

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


    Proof.
      intros t1 t2.

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

1 subgoal (ID 40)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  jobs : seq Job
  H_unit_service : unit_service_proc_model PState
  t1, t2 : instant
  ============================
  service_of_jobs sched P jobs t1 t2 <= workload_of_jobs P jobs

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


      apply leq_sum; intros j _.

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

1 subgoal (ID 44)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  jobs : seq Job
  H_unit_service : unit_service_proc_model PState
  t1, t2 : instant
  j : Job
  ============================
  service_during sched j t1 t2 <= job_cost j

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


        by apply cumulative_service_le_job_cost.

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

No more subgoals.

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


    Qed.

  End ServiceBoundedByWorkload.

In this section we prove a few facts about splitting the total service of a set of jobs.
  Section ServiceCat.

We show that the total service of jobs released in a time interval [t1,t2) during [t1,t2) is equal to the sum of: (1) the total service of jobs released in time interval [t1, t) during time [t1, t) (2) the total service of jobs released in time interval [t1, t) during time [t, t2) and (3) the total service of jobs released in time interval [t, t2) during time [t, t2).
    Lemma service_of_jobs_cat_scheduling_interval :
       t1 t2 t,
        t1 t t2
        service_of_jobs sched P (arrivals_between t1 t2) t1 t2
        = service_of_jobs sched P (arrivals_between t1 t) t1 t
          + service_of_jobs sched P (arrivals_between t1 t) t t2
          + service_of_jobs sched P (arrivals_between t t2) t t2.

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

1 subgoal (ID 45)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  ============================
  forall t1 t2 t : nat,
  t1 <= t <= t2 ->
  service_of_jobs sched P (arrivals_between t1 t2) t1 t2 =
  service_of_jobs sched P (arrivals_between t1 t) t1 t +
  service_of_jobs sched P (arrivals_between t1 t) t t2 +
  service_of_jobs sched P (arrivals_between t t2) t t2

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


    Proof.
      movet1 t2 t /andP [GEt LEt].

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

1 subgoal (ID 87)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  t1, t2, t : nat
  GEt : t1 <= t
  LEt : t <= t2
  ============================
  service_of_jobs sched P (arrivals_between t1 t2) t1 t2 =
  service_of_jobs sched P (arrivals_between t1 t) t1 t +
  service_of_jobs sched P (arrivals_between t1 t) t t2 +
  service_of_jobs sched P (arrivals_between t t2) t t2

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


      rewrite /arrivals_between (arrivals_between_cat _ _ t) //.

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

1 subgoal (ID 103)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  t1, t2, t : nat
  GEt : t1 <= t
  LEt : t <= t2
  ============================
  service_of_jobs sched P
    (arrival_sequence.arrivals_between arr_seq t1 t ++
     arrival_sequence.arrivals_between arr_seq t t2) t1 t2 =
  service_of_jobs sched P (arrival_sequence.arrivals_between arr_seq t1 t) t1
    t +
  service_of_jobs sched P (arrival_sequence.arrivals_between arr_seq t1 t) t
    t2 +
  service_of_jobs sched P (arrival_sequence.arrivals_between arr_seq t t2) t
    t2

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


      rewrite {1}/service_of_jobs big_cat //=.

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

1 subgoal (ID 144)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  t1, t2, t : nat
  GEt : t1 <= t
  LEt : t <= t2
  ============================
  \sum_(i <- arrival_sequence.arrivals_between arr_seq t1 t | 
  P i) service_during sched i t1 t2 +
  \sum_(i <- arrival_sequence.arrivals_between arr_seq t t2 | 
  P i) service_during sched i t1 t2 =
  service_of_jobs sched P (arrival_sequence.arrivals_between arr_seq t1 t) t1
    t +
  service_of_jobs sched P (arrival_sequence.arrivals_between arr_seq t1 t) t
    t2 +
  service_of_jobs sched P (arrival_sequence.arrivals_between arr_seq t t2) t
    t2

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


      rewrite exchange_big //= (@big_cat_nat _ _ _ t) //=;
              rewrite [in X in X + _ + _]exchange_big //= [in X in _ + X + _]exchange_big //=.

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

1 subgoal (ID 308)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  t1, t2, t : nat
  GEt : t1 <= t
  LEt : t <= t2
  ============================
  \sum_(j <- arrival_sequence.arrivals_between arr_seq t1 t | 
  P j) \sum_(t1 <= i < t) service_at sched j i +
  \sum_(j <- arrival_sequence.arrivals_between arr_seq t1 t | 
  P j) \sum_(t <= i < t2) service_at sched j i +
  \sum_(i <- arrival_sequence.arrivals_between arr_seq t t2 | 
  P i) service_during sched i t1 t2 =
  service_of_jobs sched P (arrival_sequence.arrivals_between arr_seq t1 t) t1
    t +
  service_of_jobs sched P (arrival_sequence.arrivals_between arr_seq t1 t) t
    t2 +
  service_of_jobs sched P (arrival_sequence.arrivals_between arr_seq t t2) t
    t2

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


      apply/eqP; rewrite -!addnA eqn_add2l eqn_add2l.

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

1 subgoal (ID 413)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  t1, t2, t : nat
  GEt : t1 <= t
  LEt : t <= t2
  ============================
  \sum_(i <- arrival_sequence.arrivals_between arr_seq t t2 | 
  P i) service_during sched i t1 t2 ==
  service_of_jobs sched P (arrival_sequence.arrivals_between arr_seq t t2) t
    t2

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


      rewrite exchange_big //= (@big_cat_nat _ _ _ t) //= [in X in _ + X]exchange_big //=.

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

1 subgoal (ID 511)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  t1, t2, t : nat
  GEt : t1 <= t
  LEt : t <= t2
  ============================
  \sum_(t1 <= i < t)
     \sum_(i0 <- arrival_sequence.arrivals_between arr_seq t t2 | 
     P i0) service_at sched i0 i +
  \sum_(j <- arrival_sequence.arrivals_between arr_seq t t2 | 
  P j) \sum_(t <= i < t2) service_at sched j i ==
  service_of_jobs sched P (arrival_sequence.arrivals_between arr_seq t t2) t
    t2

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


      rewrite -[service_of_jobs _ _ _ _ _]add0n /service_of_jobs eqn_add2r.

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

1 subgoal (ID 562)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  t1, t2, t : nat
  GEt : t1 <= t
  LEt : t <= t2
  ============================
  \sum_(t1 <= i < t)
     \sum_(i0 <- arrival_sequence.arrivals_between arr_seq t t2 | 
     P i0) service_at sched i0 i == 0

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


      rewrite big_nat_cond big1 //.

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

1 subgoal (ID 582)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  t1, t2, t : nat
  GEt : t1 <= t
  LEt : t <= t2
  ============================
  forall i : nat,
  (t1 <= i < t) && true ->
  \sum_(i0 <- arrival_sequence.arrivals_between arr_seq t t2 | 
  P i0) service_at sched i0 i = 0

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


      movex /andP [/andP [GEi LTi] _].

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

1 subgoal (ID 684)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  t1, t2, t : nat
  GEt : t1 <= t
  LEt : t <= t2
  x : nat
  GEi : t1 <= x
  LTi : x < t
  ============================
  \sum_(i <- arrival_sequence.arrivals_between arr_seq t t2 | 
  P i) service_at sched i x = 0

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


      rewrite big_seq_cond big1 //.

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

1 subgoal (ID 704)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  t1, t2, t : nat
  GEt : t1 <= t
  LEt : t <= t2
  x : nat
  GEi : t1 <= x
  LTi : x < t
  ============================
  forall i : Job,
  (i \in arrival_sequence.arrivals_between arr_seq t t2) && P i ->
  service_at sched i x = 0

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


      movej /andP [ARR Ps].

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

1 subgoal (ID 769)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  t1, t2, t : nat
  GEt : t1 <= t
  LEt : t <= t2
  x : nat
  GEi : t1 <= x
  LTi : x < t
  j : Job
  ARR : j \in arrival_sequence.arrivals_between arr_seq t t2
  Ps : P j
  ============================
  service_at sched j x = 0

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


      apply service_before_job_arrival_zero with H0; auto.

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

1 subgoal (ID 774)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  t1, t2, t : nat
  GEt : t1 <= t
  LEt : t <= t2
  x : nat
  GEi : t1 <= x
  LTi : x < t
  j : Job
  ARR : j \in arrival_sequence.arrivals_between arr_seq t t2
  Ps : P j
  ============================
  x < job_arrival j

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


      eapply in_arrivals_implies_arrived_between in ARR; eauto 2.

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

1 subgoal (ID 803)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  t1, t2, t : nat
  GEt : t1 <= t
  LEt : t <= t2
  x : nat
  GEi : t1 <= x
  LTi : x < t
  j : Job
  ARR : arrived_between j t t2
  Ps : P j
  ============================
  x < job_arrival j

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


        by move: ARR ⇒ /andP [N1 N2]; apply leq_trans with t.

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

No more subgoals.

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


    Qed.

We show that the total service of jobs released in a time interval [t1,t2) during [t,t2) is equal to the sum of: (1) the total service of jobs released in a time interval [t1,t) during [t,t2) and (2) the total service of jobs released in a time interval [t,t2) during [t,t2).
    Lemma service_of_jobs_cat_arrival_interval :
       t1 t2 t,
        t1 t t2
        service_of_jobs sched P (arrivals_between t1 t2) t t2 =
        service_of_jobs sched P (arrivals_between t1 t) t t2 +
        service_of_jobs sched P (arrivals_between t t2) t t2.

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

1 subgoal (ID 59)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  ============================
  forall t1 t2 t : nat,
  t1 <= t <= t2 ->
  service_of_jobs sched P (arrivals_between t1 t2) t t2 =
  service_of_jobs sched P (arrivals_between t1 t) t t2 +
  service_of_jobs sched P (arrivals_between t t2) t t2

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


    Proof.
      movet1 t2 t /andP [GEt LEt].

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

1 subgoal (ID 101)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  t1, t2, t : nat
  GEt : t1 <= t
  LEt : t <= t2
  ============================
  service_of_jobs sched P (arrivals_between t1 t2) t t2 =
  service_of_jobs sched P (arrivals_between t1 t) t t2 +
  service_of_jobs sched P (arrivals_between t t2) t t2

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


      apply/eqP;rewrite eq_sym; apply/eqP.

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

1 focused subgoal
(shelved: 1) (ID 189)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  t1, t2, t : nat
  GEt : t1 <= t
  LEt : t <= t2
  ============================
  service_of_jobs sched P (arrivals_between t1 t) t t2 +
  service_of_jobs sched P (arrivals_between t t2) t t2 =
  service_of_jobs sched P (arrivals_between t1 t2) t t2

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


      rewrite /arrivals_between [in X in _ = X](arrivals_between_cat _ _ t) //.

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

1 subgoal (ID 212)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  PState : Type
  H2 : ProcessorState Job PState
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule PState
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  t1, t2, t : nat
  GEt : t1 <= t
  LEt : t <= t2
  ============================
  service_of_jobs sched P (arrival_sequence.arrivals_between arr_seq t1 t) t
    t2 +
  service_of_jobs sched P (arrival_sequence.arrivals_between arr_seq t t2) t
    t2 =
  service_of_jobs sched P
    (arrival_sequence.arrivals_between arr_seq t1 t ++
     arrival_sequence.arrivals_between arr_seq t t2) t t2

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


        by rewrite {3}/service_of_jobs -big_cat //=.

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

No more subgoals.

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


    Qed.

  End ServiceCat.

End GenericModelLemmas.

In this section, we prove some properties about service of sets of jobs for ideal uni-processor model.
Section IdealModelLemmas.

Consider any type of tasks ...
  Context {Task : TaskType}.

... and any type of jobs associated with these tasks.
  Context {Job : JobType}.
  Context `{JobTask Job Task}.
  Context `{JobArrival Job}.
  Context `{JobCost Job}.

Consider any arrival sequence with consistent arrivals.
Next, consider any ideal uni-processor schedule of this arrival sequence ...
... where jobs do not execute before their arrival or after completion.
Let P be any predicate over jobs.
  Variable P : Job bool.

Let's define some local names for clarity.
We prove that if the total service within some time interval [t1,t2) is smaller than [t2-t1], then there is an idle time instant t ∈ [[t1,t2)].
  Lemma low_service_implies_existence_of_idle_time :
     t1 t2,
      service_of_jobs sched predT (arrivals_between 0 t2) t1 t2 < t2 - t1
       t, t1 t < t2 is_idle sched t.

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

1 subgoal (ID 45)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  ============================
  forall t1 t2 : instant,
  service_of_jobs sched predT (arrivals_between 0 t2) t1 t2 < t2 - t1 ->
  exists t : nat, t1 <= t < t2 /\ is_idle sched t

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


  Proof.
    intros ? ? SERV.

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

1 subgoal (ID 48)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  SERV : service_of_jobs sched predT (arrivals_between 0 t2) t1 t2 < t2 - t1
  ============================
  exists t : nat, t1 <= t < t2 /\ is_idle sched t

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


    destruct (t1 t2) eqn:LE; last first.

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

2 subgoals (ID 59)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  SERV : service_of_jobs sched predT (arrivals_between 0 t2) t1 t2 < t2 - t1
  LE : (t1 <= t2) = false
  ============================
  exists t : nat, t1 <= t < t2 /\ is_idle sched t

subgoal 2 (ID 58) is:
 exists t : nat, t1 <= t < t2 /\ is_idle sched t

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


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

1 subgoal (ID 59)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  SERV : service_of_jobs sched predT (arrivals_between 0 t2) t1 t2 < t2 - t1
  LE : (t1 <= t2) = false
  ============================
  exists t : nat, t1 <= t < t2 /\ is_idle sched t

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


move: LE ⇒ /negP/negP; rewrite -ltnNge.

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

1 subgoal (ID 140)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  SERV : service_of_jobs sched predT (arrivals_between 0 t2) t1 t2 < t2 - t1
  ============================
  t2 < t1 -> exists t : nat, t1 <= t < t2 /\ is_idle sched t

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


      moveLT; apply ltnW in LT; rewrite -subn_eq0 in LT.

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

1 subgoal (ID 172)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  SERV : service_of_jobs sched predT (arrivals_between 0 t2) t1 t2 < t2 - t1
  LT : t2 - t1 == 0
  ============================
  exists t : nat, t1 <= t < t2 /\ is_idle sched t

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


        by rewrite (eqbool_to_eqprop LT) ltn0 in SERV.

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

1 subgoal

subgoal 1 (ID 58) is:
 exists t : nat, t1 <= t < t2 /\ is_idle sched t

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


    }

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

1 subgoal (ID 58)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  SERV : service_of_jobs sched predT (arrivals_between 0 t2) t1 t2 < t2 - t1
  LE : (t1 <= t2) = true
  ============================
  exists t : nat, t1 <= t < t2 /\ is_idle sched t

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


    have EX: δ, t2 = t1 + δ.

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

2 subgoals (ID 238)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  SERV : service_of_jobs sched predT (arrivals_between 0 t2) t1 t2 < t2 - t1
  LE : (t1 <= t2) = true
  ============================
  exists δ : nat, t2 = t1 + δ

subgoal 2 (ID 240) is:
 exists t : nat, t1 <= t < t2 /\ is_idle sched t

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


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

1 subgoal (ID 238)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  SERV : service_of_jobs sched predT (arrivals_between 0 t2) t1 t2 < t2 - t1
  LE : (t1 <= t2) = true
  ============================
  exists δ : nat, t2 = t1 + δ

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


by (t2 - t1); rewrite subnKC // ltnW.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal

subgoal 1 (ID 240) is:
 exists t : nat, t1 <= t < t2 /\ is_idle sched t

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


}

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

1 subgoal (ID 240)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  SERV : service_of_jobs sched predT (arrivals_between 0 t2) t1 t2 < t2 - t1
  LE : (t1 <= t2) = true
  EX : exists δ : nat, t2 = t1 + δ
  ============================
  exists t : nat, t1 <= t < t2 /\ is_idle sched t

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


    move: EX ⇒ [δ EQ]; subst t2; clear LE.

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

1 subgoal (ID 270)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1 : instant
  δ : nat
  SERV : service_of_jobs sched predT (arrivals_between 0 (t1 + δ)) t1
           (t1 + δ) < t1 + δ - t1
  ============================
  exists t : nat, t1 <= t < t1 + δ /\ is_idle sched t

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


    rewrite {3}[t1 + δ]addnC -addnBA // subnn addn0 in SERV.

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

1 subgoal (ID 345)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1 : instant
  δ : nat
  SERV : service_of_jobs sched predT (arrivals_between 0 (t1 + δ)) t1
           (t1 + δ) < δ
  ============================
  exists t : nat, t1 <= t < t1 + δ /\ is_idle sched t

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


    rewrite /service_of_jobs exchange_big //= in SERV.

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

1 subgoal (ID 417)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1 : instant
  δ : nat
  SERV : \sum_(t1 <= j < t1 + δ)
            \sum_(i <- arrivals_between 0 (t1 + δ)) (sched j == Some i) < δ
  ============================
  exists t : nat, t1 <= t < t1 + δ /\ is_idle sched t

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


    apply sum_le_summation_range in SERV.

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

1 subgoal (ID 419)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1 : instant
  δ : nat
  SERV : exists x : nat,
           t1 <= x < t1 + δ /\
           \sum_(i <- arrivals_between 0 (t1 + δ)) (sched x == Some i) = 0
  ============================
  exists t : nat, t1 <= t < t1 + δ /\ is_idle sched t

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


    move: SERV ⇒ [x [/andP [GEx LEx] L]].

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

1 subgoal (ID 480)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1 : instant
  δ, x : nat
  GEx : t1 <= x
  LEx : x < t1 + δ
  L : \sum_(i <- arrivals_between 0 (t1 + δ)) (sched x == Some i) = 0
  ============================
  exists t : nat, t1 <= t < t1 + δ /\ is_idle sched t

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


     x; split; first by apply/andP; split.

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

1 subgoal (ID 485)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1 : instant
  δ, x : nat
  GEx : t1 <= x
  LEx : x < t1 + δ
  L : \sum_(i <- arrivals_between 0 (t1 + δ)) (sched x == Some i) = 0
  ============================
  is_idle sched x

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


    apply/negPn; apply/negP; intros NIDLE.

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

1 subgoal (ID 576)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1 : instant
  δ, x : nat
  GEx : t1 <= x
  LEx : x < t1 + δ
  L : \sum_(i <- arrivals_between 0 (t1 + δ)) (sched x == Some i) = 0
  NIDLE : ~~ is_idle sched x
  ============================
  False

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


    unfold is_idle in NIDLE.

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

1 subgoal (ID 577)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1 : instant
  δ, x : nat
  GEx : t1 <= x
  LEx : x < t1 + δ
  L : \sum_(i <- arrivals_between 0 (t1 + δ)) (sched x == Some i) = 0
  NIDLE : sched x != None
  ============================
  False

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


    destruct(sched x) eqn:SCHED; last by done.

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

1 subgoal (ID 596)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1 : instant
  δ, x : nat
  GEx : t1 <= x
  LEx : x < t1 + δ
  s : Job
  SCHED : sched x = Some s
  L : \sum_(i <- arrivals_between 0 (t1 + δ)) (Some s == Some i) = 0
  NIDLE : Some s != None
  ============================
  False

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


    move: SCHED ⇒ /eqP SCHED; clear NIDLE; rewrite -scheduled_at_def/= in SCHED.

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

1 subgoal (ID 724)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1 : instant
  δ, x : nat
  GEx : t1 <= x
  LEx : x < t1 + δ
  s : Job
  L : \sum_(i <- arrivals_between 0 (t1 + δ)) (Some s == Some i) = 0
  SCHED : scheduled_at sched s x
  ============================
  False

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


    move: L ⇒ /eqP; rewrite -sum_nat_eq0_nat; move ⇒ /allP L.

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

1 subgoal (ID 821)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1 : instant
  δ, x : nat
  GEx : t1 <= x
  LEx : x < t1 + δ
  s : Job
  SCHED : scheduled_at sched s x
  L : {in arrivals_between 0 (t1 + δ),
        forall x : Job, (Some s == Some x) == 0}
  ============================
  False

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


    specialize (L s); feed L.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 824)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1 : instant
  δ, x : nat
  GEx : t1 <= x
  LEx : x < t1 + δ
  s : Job
  SCHED : scheduled_at sched s x
  L : s \in arrivals_between 0 (t1 + δ) ->
      (fun x : Job => is_true ((Some s == Some x) == 0)) s
  ============================
  s \in arrivals_between 0 (t1 + δ)

subgoal 2 (ID 829) is:
 False

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


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

1 subgoal (ID 824)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1 : instant
  δ, x : nat
  GEx : t1 <= x
  LEx : x < t1 + δ
  s : Job
  SCHED : scheduled_at sched s x
  L : s \in arrivals_between 0 (t1 + δ) ->
      (fun x : Job => is_true ((Some s == Some x) == 0)) s
  ============================
  s \in arrivals_between 0 (t1 + δ)

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


unfold arrivals_between.

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

1 subgoal (ID 830)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1 : instant
  δ, x : nat
  GEx : t1 <= x
  LEx : x < t1 + δ
  s : Job
  SCHED : scheduled_at sched s x
  L : s \in arrivals_between 0 (t1 + δ) ->
      (fun x : Job => is_true ((Some s == Some x) == 0)) s
  ============================
  s \in arrival_sequence.arrivals_between arr_seq 0 (t1 + δ)

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


      eapply arrived_between_implies_in_arrivals; eauto 2.

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

1 subgoal (ID 835)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1 : instant
  δ, x : nat
  GEx : t1 <= x
  LEx : x < t1 + δ
  s : Job
  SCHED : scheduled_at sched s x
  L : s \in arrivals_between 0 (t1 + δ) ->
      (fun x : Job => is_true ((Some s == Some x) == 0)) s
  ============================
  arrived_between s 0 (t1 + δ)

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


        by apply H_jobs_must_arrive_to_execute in SCHED; apply leq_ltn_trans with x.

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

1 subgoal

subgoal 1 (ID 829) is:
 False

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


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

1 subgoal (ID 829)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1 : instant
  δ, x : nat
  GEx : t1 <= x
  LEx : x < t1 + δ
  s : Job
  SCHED : scheduled_at sched s x
  L : (fun x : Job => is_true ((Some s == Some x) == 0)) s
  ============================
  False

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



      by move: L; simpl;rewrite eqb0; move ⇒ /eqP EQ.

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

No more subgoals.

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


  Qed.

In this section, we prove that the service received by any set of jobs is upper-bounded by the corresponding interval length.
Let [jobs] denote any (finite) set of jobs.
    Variable jobs : seq Job.

Assume that the sequence of [jobs] is a set.
    Hypothesis H_no_duplicate_jobs : uniq jobs.

We prove that the overall service of [jobs] at a single time instant is at most [1].
    Lemma service_of_jobs_le_1:
       t, \sum_(j <- jobs | P j) service_at sched j t 1.

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

1 subgoal (ID 56)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  ============================
  forall t : instant, \sum_(j <- jobs | P j) service_at sched j t <= 1

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


    Proof.
      intros t.

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

1 subgoal (ID 57)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  ============================
  \sum_(j <- jobs | P j) service_at sched j t <= 1

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


      case SCHED: (sched t) ⇒ [j | ]; simpl.

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

2 subgoals (ID 107)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ============================
  \sum_(j0 <- jobs | P j0) (sched t == Some j0) <= 1

subgoal 2 (ID 108) is:
 \sum_(j <- jobs | P j) (sched t == Some j) <= 1

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


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

1 subgoal (ID 107)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ============================
  \sum_(j0 <- jobs | P j0) (sched t == Some j0) <= 1

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


case ARR: (j \in jobs).

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

2 subgoals (ID 141)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = true
  ============================
  \sum_(j0 <- jobs | P j0) (sched t == Some j0) <= 1

subgoal 2 (ID 162) is:
 \sum_(j0 <- jobs | P j0) (sched t == Some j0) <= 1

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


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

1 subgoal (ID 141)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = true
  ============================
  \sum_(j0 <- jobs | P j0) (sched t == Some j0) <= 1

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


rewrite (big_rem j) //=; simpl.

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

1 subgoal (ID 202)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = true
  ============================
  (if P j then sched t == Some j else 0) +
  \sum_(y <- rem (T:=Job) j jobs | P y) (sched t == Some y) <= 1

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


          rewrite /service_at /scheduled_at SCHED; simpl.

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

1 subgoal (ID 219)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = true
  ============================
  (if P j then Some j == Some j else 0) +
  \sum_(y <- rem (T:=Job) j jobs | P y) (Some j == Some y) <= 1

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


          rewrite -[1]addn0 leq_add //.

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

2 subgoals (ID 233)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = true
  ============================
  (if P j then Some j == Some j else 0) <= 1

subgoal 2 (ID 234) is:
 \sum_(y <- rem (T:=Job) j jobs | P y) (Some j == Some y) <= 0

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


          - by rewrite eq_refl; case (P j).

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

1 subgoal (ID 234)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = true
  ============================
  \sum_(y <- rem (T:=Job) j jobs | P y) (Some j == Some y) <= 0

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


          - rewrite leqn0 big1_seq; first by done.

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

1 subgoal (ID 296)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = true
  ============================
  forall i : Job,
  P i && (i \in rem (T:=Job) j jobs) -> (Some j == Some i) = 0

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


            movej' /andP [_ ARRj'].

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

1 subgoal (ID 336)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = true
  j' : Job
  ARRj' : j' \in rem (T:=Job) j jobs
  ============================
  (Some j == Some j') = 0

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


            apply/eqP; rewrite eqb0.

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

1 subgoal (ID 394)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = true
  j' : Job
  ARRj' : j' \in rem (T:=Job) j jobs
  ============================
  Some j != Some j'

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


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

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

1 subgoal (ID 451)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = true
  j' : Job
  ARRj' : j' \in rem (T:=Job) j jobs
  CONTR : Some j = Some j'
  ============================
  False

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


            inversion CONTR; subst j'; clear CONTR.

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

1 subgoal (ID 477)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = true
  ARRj' : j \in rem (T:=Job) j jobs
  ============================
  False

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


            rewrite rem_filter in ARRj'; last by done.

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

1 subgoal (ID 512)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = true
  ARRj' : j \in [seq x <- jobs | predC1 j x]
  ============================
  False

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


            move: ARRj'; rewrite mem_filter; move ⇒ /andP [/negP CONTR _].

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

1 subgoal (ID 605)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = true
  CONTR : ~ j == j
  ============================
  False

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


              by apply: CONTR.

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

2 subgoals

subgoal 1 (ID 162) is:
 \sum_(j0 <- jobs | P j0) (sched t == Some j0) <= 1
subgoal 2 (ID 108) is:
 \sum_(j <- jobs | P j) (sched t == Some j) <= 1

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


        }

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

1 subgoal (ID 162)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = false
  ============================
  \sum_(j0 <- jobs | P j0) (sched t == Some j0) <= 1

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


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

1 subgoal (ID 162)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = false
  ============================
  \sum_(j0 <- jobs | P j0) (sched t == Some j0) <= 1

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


apply leq_trans with 0; last by done.

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

1 subgoal (ID 612)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = false
  ============================
  \sum_(j0 <- jobs | P j0) (sched t == Some j0) <= 0

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


          rewrite leqn0 big1_seq; first by done.

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

1 subgoal (ID 627)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = false
  ============================
  forall i : Job, P i && (i \in jobs) -> (sched t == Some i) = 0

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


          movej' /andP [_ ARRj'].

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

1 subgoal (ID 667)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = false
  j' : Job
  ARRj' : j' \in jobs
  ============================
  (sched t == Some j') = 0

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


          apply/eqP; rewrite eqb0.

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

1 subgoal (ID 725)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = false
  j' : Job
  ARRj' : j' \in jobs
  ============================
  sched t != Some j'

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


          rewrite /scheduled_at SCHED.

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

1 subgoal (ID 734)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = false
  j' : Job
  ARRj' : j' \in jobs
  ============================
  Some j != Some j'

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


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

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

1 subgoal (ID 791)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = false
  j' : Job
  ARRj' : j' \in jobs
  CONTR : Some j = Some j'
  ============================
  False

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


          inversion CONTR; clear CONTR.

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

1 subgoal (ID 809)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  j : Job
  SCHED : sched t = Some j
  ARR : (j \in jobs) = false
  j' : Job
  ARRj' : j' \in jobs
  H3 : j = j'
  ============================
  False

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


            by subst j'; rewrite ARR in ARRj'.

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

1 subgoal

subgoal 1 (ID 108) is:
 \sum_(j <- jobs | P j) (sched t == Some j) <= 1

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


        }

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

1 subgoal

subgoal 1 (ID 108) is:
 \sum_(j <- jobs | P j) (sched t == Some j) <= 1

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


      }

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

1 subgoal (ID 108)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  SCHED : sched t = None
  ============================
  \sum_(j <- jobs | P j) (sched t == Some j) <= 1

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


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

1 subgoal (ID 108)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  SCHED : sched t = None
  ============================
  \sum_(j <- jobs | P j) (sched t == Some j) <= 1

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


apply leq_trans with 0; last by done.

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

1 subgoal (ID 872)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  SCHED : sched t = None
  ============================
  \sum_(j <- jobs | P j) (sched t == Some j) <= 0

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


        rewrite leqn0 big1_seq; first by done.

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

1 subgoal (ID 887)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  SCHED : sched t = None
  ============================
  forall i : Job, P i && (i \in jobs) -> (sched t == Some i) = 0

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


        movej' /andP [_ ARRj'].

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

1 subgoal (ID 927)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  SCHED : sched t = None
  j' : Job
  ARRj' : j' \in jobs
  ============================
  (sched t == Some j') = 0

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


          by rewrite /service_at /scheduled_at SCHED.

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

No more subgoals.

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


      }

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

No more subgoals.

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


    Qed.

Then, we prove that the service received by those jobs is no larger than their workload.
    Corollary service_of_jobs_le_length_of_interval:
       (t : instant) (Δ : duration),
        service_of_jobs sched P jobs t (t + Δ) Δ.

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

1 subgoal (ID 60)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  ============================
  forall (t : instant) (Δ : duration),
  service_of_jobs sched P jobs t (t + Δ) <= Δ

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


    Proof.
      intros.

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

1 subgoal (ID 62)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  Δ : duration
  ============================
  service_of_jobs sched P jobs t (t + Δ) <= Δ

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


      have EQ: \sum_(t x < t + Δ) 1 = Δ.

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

2 subgoals (ID 68)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  Δ : duration
  ============================
  \sum_(t <= x < t + Δ) 1 = Δ

subgoal 2 (ID 70) is:
 service_of_jobs sched P jobs t (t + Δ) <= Δ

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


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

1 subgoal (ID 68)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  Δ : duration
  ============================
  \sum_(t <= x < t + Δ) 1 = Δ

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


by rewrite big_const_nat iter_addn mul1n addn0 -{2}[t]addn0 subnDl subn0.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal

subgoal 1 (ID 70) is:
 service_of_jobs sched P jobs t (t + Δ) <= Δ

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


}

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

1 subgoal (ID 70)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  Δ : duration
  EQ : \sum_(t <= x < t + Δ) 1 = Δ
  ============================
  service_of_jobs sched P jobs t (t + Δ) <= Δ

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


      rewrite -{2}EQ {EQ}.

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

1 subgoal (ID 111)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  Δ : duration
  ============================
  service_of_jobs sched P jobs t (t + Δ) <= \sum_(t <= x < t + Δ) 1

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


      rewrite /service_of_jobs/service_during/service_at exchange_big //=.

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

1 subgoal (ID 145)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  Δ : duration
  ============================
  \sum_(t <= j < t + Δ) \sum_(i <- jobs | P i) service_in i (sched j) <=
  \sum_(t <= x < t + Δ) 1

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


      rewrite leq_sum //.

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

1 subgoal (ID 175)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  Δ : duration
  ============================
  forall i : nat,
  true -> \sum_(i0 <- jobs | P i0) service_in i0 (sched i) <= 1

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


      movet' _.

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

1 subgoal (ID 203)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t : instant
  Δ : duration
  t' : nat
  ============================
  \sum_(i <- jobs | P i) service_in i (sched t') <= 1

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


        by apply service_of_jobs_le_1.

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

No more subgoals.

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


    Qed.

Similar lemma holds for two instants.
    Corollary service_of_jobs_le_length_of_interval':
       (t1 t2 : instant),
        service_of_jobs sched P jobs t1 t2 t2 - t1.

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

1 subgoal (ID 64)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  ============================
  forall t1 t2 : instant, service_of_jobs sched P jobs t1 t2 <= t2 - t1

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


    Proof.
      intros.

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

1 subgoal (ID 66)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t1, t2 : instant
  ============================
  service_of_jobs sched P jobs t1 t2 <= t2 - t1

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


      have <-: \sum_(t1 x < t2) 1 = t2 - t1.

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

2 subgoals (ID 72)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t1, t2 : instant
  ============================
  \sum_(t1 <= x < t2) 1 = t2 - t1

subgoal 2 (ID 77) is:
 service_of_jobs sched P jobs t1 t2 <= \sum_(t1 <= x < t2) 1

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


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

1 subgoal (ID 72)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t1, t2 : instant
  ============================
  \sum_(t1 <= x < t2) 1 = t2 - t1

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


by rewrite big_const_nat iter_addn mul1n addn0.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal

subgoal 1 (ID 77) is:
 service_of_jobs sched P jobs t1 t2 <= \sum_(t1 <= x < t2) 1

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


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

1 subgoal (ID 77)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t1, t2 : instant
  ============================
  service_of_jobs sched P jobs t1 t2 <= \sum_(t1 <= x < t2) 1

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



      rewrite /service_of_jobs exchange_big //=.

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

1 subgoal (ID 120)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t1, t2 : instant
  ============================
  \sum_(t1 <= j < t2) \sum_(i <- jobs | P i) (sched j == Some i) <=
  \sum_(t1 <= x < t2) 1

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


      rewrite leq_sum //.

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

1 subgoal (ID 150)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t1, t2 : instant
  ============================
  forall i : nat, true -> \sum_(i0 <- jobs | P i0) (sched i == Some i0) <= 1

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


      movet' _.

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

1 subgoal (ID 178)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t1, t2 : instant
  t' : nat
  ============================
  \sum_(i <- jobs | P i) (sched t' == Some i) <= 1

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


      have SE := service_of_jobs_le_1 t'.

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

1 subgoal (ID 183)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t1, t2 : instant
  t' : nat
  SE : \sum_(j <- jobs | P j) service_at sched j t' <= 1
  ============================
  \sum_(i <- jobs | P i) (sched t' == Some i) <= 1

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


      eapply leq_trans; last by eassumption.

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

1 subgoal (ID 185)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  jobs : seq Job
  H_no_duplicate_jobs : uniq jobs
  t1, t2 : instant
  t' : nat
  SE : \sum_(j <- jobs | P j) service_at sched j t' <= 1
  ============================
  \sum_(i <- jobs | P i) (sched t' == Some i) <=
  \sum_(j <- jobs | P j) service_at sched j t'

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


        by rewrite leq_sum //.

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

No more subgoals.

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


    Qed.

  End ServiceOfJobsIsBoundedByLength.

In this section, we introduce a connection between the cumulative service, cumulative workload, and completion of jobs.
Consider an arbitrary time interval [t1, t2).
    Variables t1 t2 : instant.

Let jobs be a set of all jobs arrived during [t1, t2).
    Let jobs := arrivals_between t1 t2.

Next, we consider some time instant [t_compl].
    Variable t_compl : instant.

And state the proposition that all jobs are completed by time [t_compl]. Next we show that this proposition is equivalent to the fact that [workload of jobs = service of jobs].
    Let all_jobs_completed_by t_compl :=
       j, j \in jobs P j completed_by j t_compl.

First, we prove that if the workload of [jobs] is equal to the service of [jobs], then any job in [jobs] is completed by time [t_compl].
    Lemma workload_eq_service_impl_all_jobs_have_completed:
      workload_of_jobs P jobs =
      service_of_jobs sched P jobs t1 t_compl
      all_jobs_completed_by t_compl.

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

1 subgoal (ID 59)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  ============================
  workload_of_jobs P jobs = service_of_jobs sched P jobs t1 t_compl ->
  all_jobs_completed_by t_compl

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


    Proof.
      unfold jobs; intros EQ j ARR Pj; move: (ARR) ⇒ ARR2.

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

1 subgoal (ID 67)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  EQ : workload_of_jobs P (arrivals_between t1 t2) =
       service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl
  j : Job
  ARR : j \in jobs
  Pj : P j
  ARR2 : j \in jobs
  ============================
  completed_by j t_compl

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


      apply in_arrivals_implies_arrived_between in ARR; last by done.

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

1 subgoal (ID 71)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  EQ : workload_of_jobs P (arrivals_between t1 t2) =
       service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl
  j : Job
  ARR : arrived_between j t1 t2
  Pj : P j
  ARR2 : j \in jobs
  ============================
  completed_by j t_compl

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


      move: ARR ⇒ /andP [T1 T2].

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

1 subgoal (ID 113)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  EQ : workload_of_jobs P (arrivals_between t1 t2) =
       service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  ============================
  completed_by j t_compl

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


      have F1: a b, (a < b) || (a b).

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

2 subgoals (ID 116)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  EQ : workload_of_jobs P (arrivals_between t1 t2) =
       service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  ============================
  forall a b : nat, (a < b) || (b <= a)

subgoal 2 (ID 118) is:
 completed_by j t_compl

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


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

1 subgoal (ID 116)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  EQ : workload_of_jobs P (arrivals_between t1 t2) =
       service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  ============================
  forall a b : nat, (a < b) || (b <= a)

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


by intros; destruct (a < b) eqn:EQU; apply/orP;
          [by left |right]; apply negbT in EQU; rewrite leqNgt.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal

subgoal 1 (ID 118) is:
 completed_by j t_compl

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


}

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

1 subgoal (ID 118)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  EQ : workload_of_jobs P (arrivals_between t1 t2) =
       service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  F1 : forall a b : nat, (a < b) || (b <= a)
  ============================
  completed_by j t_compl

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


      move: (F1 t_compl t1) ⇒ /orP [LT | GT].

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

2 subgoals (ID 234)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  EQ : workload_of_jobs P (arrivals_between t1 t2) =
       service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  F1 : forall a b : nat, (a < b) || (b <= a)
  LT : t_compl < t1
  ============================
  completed_by j t_compl

subgoal 2 (ID 235) is:
 completed_by j t_compl

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


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

1 subgoal (ID 234)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  EQ : workload_of_jobs P (arrivals_between t1 t2) =
       service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  F1 : forall a b : nat, (a < b) || (b <= a)
  LT : t_compl < t1
  ============================
  completed_by j t_compl

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


rewrite /service_of_jobs /service_during in EQ.

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

1 subgoal (ID 283)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  F1 : forall a b : nat, (a < b) || (b <= a)
  LT : t_compl < t1
  EQ : workload_of_jobs P (arrivals_between t1 t2) =
       \sum_(j <- arrivals_between t1 t2 | P j)
          \sum_(t1 <= t < t_compl) service_at sched j t
  ============================
  completed_by j t_compl

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


        rewrite exchange_big big_geq //= in EQ; last by rewrite ltnW.

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

1 subgoal (ID 386)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  F1 : forall a b : nat, (a < b) || (b <= a)
  LT : t_compl < t1
  EQ : workload_of_jobs P (arrivals_between t1 t2) = 0
  ============================
  completed_by j t_compl

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


        rewrite /workload_of_jobs in EQ.

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

1 subgoal (ID 460)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  F1 : forall a b : nat, (a < b) || (b <= a)
  LT : t_compl < t1
  EQ : \sum_(j <- arrivals_between t1 t2 | P j) job_cost j = 0
  ============================
  completed_by j t_compl

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


        rewrite (big_rem j) ?Pj //= in EQ.

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

1 subgoal (ID 537)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  F1 : forall a b : nat, (a < b) || (b <= a)
  LT : t_compl < t1
  EQ : job_cost j +
       \sum_(y <- rem (T:=Job) j (arrivals_between t1 t2) | P y) job_cost y =
       0
  ============================
  completed_by j t_compl

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


        move: EQ ⇒ /eqP; rewrite addn_eq0; move ⇒ /andP [CZ _].

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

1 subgoal (ID 639)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  F1 : forall a b : nat, (a < b) || (b <= a)
  LT : t_compl < t1
  CZ : job_cost j == 0
  ============================
  completed_by j t_compl

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


        unfold completed_by, service.completed_by.

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

1 subgoal (ID 640)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  F1 : forall a b : nat, (a < b) || (b <= a)
  LT : t_compl < t1
  CZ : job_cost j == 0
  ============================
  job_cost j <= service sched j t_compl

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


          by move: CZ ⇒ /eqP CZ; rewrite CZ.

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

1 subgoal

subgoal 1 (ID 235) is:
 completed_by j t_compl

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


      }

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

1 subgoal (ID 235)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  EQ : workload_of_jobs P (arrivals_between t1 t2) =
       service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  F1 : forall a b : nat, (a < b) || (b <= a)
  GT : t1 <= t_compl
  ============================
  completed_by j t_compl

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


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

1 subgoal (ID 235)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  EQ : workload_of_jobs P (arrivals_between t1 t2) =
       service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  F1 : forall a b : nat, (a < b) || (b <= a)
  GT : t1 <= t_compl
  ============================
  completed_by j t_compl

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


unfold workload_of_jobs, service_of_jobs in EQ; unfold completed_by, service.completed_by.

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

1 subgoal (ID 679)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  EQ : \sum_(j <- arrivals_between t1 t2 | P j) job_cost j =
       \sum_(j <- arrivals_between t1 t2 | P j)
          service_during sched j t1 t_compl
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  F1 : forall a b : nat, (a < b) || (b <= a)
  GT : t1 <= t_compl
  ============================
  job_cost j <= service sched j t_compl

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


        rewrite /service -(service_during_cat _ _ _ t1); last by apply/andP; split.

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

1 subgoal (ID 708)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  EQ : \sum_(j <- arrivals_between t1 t2 | P j) job_cost j =
       \sum_(j <- arrivals_between t1 t2 | P j)
          service_during sched j t1 t_compl
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  F1 : forall a b : nat, (a < b) || (b <= a)
  GT : t1 <= t_compl
  ============================
  job_cost j <=
  service_during sched j 0 t1 + service_during sched j t1 t_compl

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


        rewrite cumulative_service_before_job_arrival_zero // add0n.

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

1 subgoal (ID 779)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  EQ : \sum_(j <- arrivals_between t1 t2 | P j) job_cost j =
       \sum_(j <- arrivals_between t1 t2 | P j)
          service_during sched j t1 t_compl
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  F1 : forall a b : nat, (a < b) || (b <= a)
  GT : t1 <= t_compl
  ============================
  job_cost j <= service_during sched j t1 t_compl

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


        rewrite <- sum_majorant_eqn with (E1 := fun jservice_during sched j t1 t_compl)
                                        (xs := arrivals_between t1 t2) (P := P); try done.

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

1 subgoal (ID 787)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  EQ : \sum_(j <- arrivals_between t1 t2 | P j) job_cost j =
       \sum_(j <- arrivals_between t1 t2 | P j)
          service_during sched j t1 t_compl
  j : Job
  Pj : P j
  ARR2 : j \in jobs
  T1 : t1 <= job_arrival j
  T2 : job_arrival j < t2
  F1 : forall a b : nat, (a < b) || (b <= a)
  GT : t1 <= t_compl
  ============================
  forall x : Job,
  x \in arrivals_between t1 t2 ->
  P x -> service_during sched x t1 t_compl <= job_cost x

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


        intros; apply cumulative_service_le_job_cost; auto using ideal_proc_model_provides_unit_service.

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

No more subgoals.

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


      }

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

No more subgoals.

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


    Qed.

And vice versa, the fact that any job in [jobs] is completed by time [t_compl] implies that the workload of [jobs] is equal to the service of [jobs].
    Lemma all_jobs_have_completed_impl_workload_eq_service:
      all_jobs_completed_by t_compl
      workload_of_jobs P jobs =
      service_of_jobs sched P jobs t1 t_compl.

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

1 subgoal (ID 66)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  ============================
  all_jobs_completed_by t_compl ->
  workload_of_jobs P jobs = service_of_jobs sched P jobs t1 t_compl

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


    Proof.
      unfold jobs; intros COMPL.

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

1 subgoal (ID 68)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  ============================
  workload_of_jobs P (arrivals_between t1 t2) =
  service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl

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


      have F: j t, t t1 j \in arrivals_between t1 t2 service_during sched j 0 t = 0.

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

2 subgoals (ID 80)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  ============================
  forall (j : Job) (t : nat),
  t <= t1 -> j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0

subgoal 2 (ID 82) is:
 workload_of_jobs P (arrivals_between t1 t2) =
 service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl

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


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

1 subgoal (ID 80)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  ============================
  forall (j : Job) (t : nat),
  t <= t1 -> j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0

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


intros j t LE ARR.

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

1 subgoal (ID 86)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  j : Job
  t : nat
  LE : t <= t1
  ARR : j \in arrivals_between t1 t2
  ============================
  service_during sched j 0 t = 0

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


        eapply in_arrivals_implies_arrived_between in ARR; eauto 2; move: ARR ⇒ /andP [GE LT].

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

1 subgoal (ID 134)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  j : Job
  t : nat
  LE : t <= t1
  GE : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  service_during sched j 0 t = 0

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


        eapply cumulative_service_before_job_arrival_zero; eauto 2.

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

1 subgoal (ID 140)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  j : Job
  t : nat
  LE : t <= t1
  GE : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  t <= job_arrival j

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


          by apply leq_trans with t1.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal

subgoal 1 (ID 82) is:
 workload_of_jobs P (arrivals_between t1 t2) =
 service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl

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


}

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

1 subgoal (ID 82)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  ============================
  workload_of_jobs P (arrivals_between t1 t2) =
  service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl

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


      destruct (t_compl t1) eqn:EQ.

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

2 subgoals (ID 161)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = true
  ============================
  workload_of_jobs P (arrivals_between t1 t2) =
  service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl

subgoal 2 (ID 162) is:
 workload_of_jobs P (arrivals_between t1 t2) =
 service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl

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


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

1 subgoal (ID 161)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = true
  ============================
  workload_of_jobs P (arrivals_between t1 t2) =
  service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl

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


unfold service_of_jobs.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 163)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = true
  ============================
  workload_of_jobs P (arrivals_between t1 t2) =
  \sum_(j <- arrivals_between t1 t2 | P j) service_during sched j t1 t_compl

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


unfold service_during.

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

1 subgoal (ID 164)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = true
  ============================
  workload_of_jobs P (arrivals_between t1 t2) =
  \sum_(j <- arrivals_between t1 t2 | P j)
     \sum_(t1 <= t < t_compl) service_at sched j t

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


        rewrite exchange_big //=.

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

1 subgoal (ID 177)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = true
  ============================
  workload_of_jobs P (arrivals_between t1 t2) =
  \sum_(t1 <= j < t_compl)
     \sum_(i <- arrivals_between t1 t2 | P i) (sched j == Some i)

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


        rewrite big_geq; last by done.

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

1 subgoal (ID 208)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = true
  ============================
  workload_of_jobs P (arrivals_between t1 t2) = 0

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


        rewrite /workload_of_jobs big1_seq //.

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

1 subgoal (ID 225)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = true
  ============================
  forall i : Job, P i && (i \in arrivals_between t1 t2) -> job_cost i = 0

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


        movej /andP [Pj ARR].

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

1 subgoal (ID 290)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = true
  j : Job
  Pj : P j
  ARR : j \in arrivals_between t1 t2
  ============================
  job_cost j = 0

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


        specialize (COMPL _ ARR Pj).

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

1 subgoal (ID 294)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  j : Job
  COMPL : completed_by j t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = true
  Pj : P j
  ARR : j \in arrivals_between t1 t2
  ============================
  job_cost j = 0

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


        rewrite <- F with (j := j) (t := t_compl); try done.

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

1 subgoal (ID 295)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  j : Job
  COMPL : completed_by j t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = true
  Pj : P j
  ARR : j \in arrivals_between t1 t2
  ============================
  job_cost j = service_during sched j 0 t_compl

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


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

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

2 subgoals (ID 403)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  j : Job
  COMPL : completed_by j t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = true
  Pj : P j
  ARR : j \in arrivals_between t1 t2
  ============================
  job_cost j <= service_during sched j 0 t_compl

subgoal 2 (ID 404) is:
 service_during sched j 0 t_compl <= job_cost j

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


        - by apply COMPL.

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

1 subgoal (ID 404)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  j : Job
  COMPL : completed_by j t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = true
  Pj : P j
  ARR : j \in arrivals_between t1 t2
  ============================
  service_during sched j 0 t_compl <= job_cost j

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


        - by apply service_at_most_cost; auto using ideal_proc_model_provides_unit_service.

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

1 subgoal

subgoal 1 (ID 162) is:
 workload_of_jobs P (arrivals_between t1 t2) =
 service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl

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


      }

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

1 subgoal (ID 162)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = false
  ============================
  workload_of_jobs P (arrivals_between t1 t2) =
  service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl

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


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

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

2 subgoals (ID 498)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = false
  ============================
  service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl <=
  workload_of_jobs P (arrivals_between t1 t2)

subgoal 2 (ID 497) is:
 workload_of_jobs P (arrivals_between t1 t2) <=
 service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl

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


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

1 subgoal (ID 498)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = false
  ============================
  service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl <=
  workload_of_jobs P (arrivals_between t1 t2)

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


by apply service_of_jobs_le_workload; auto using ideal_proc_model_provides_unit_service.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal

subgoal 1 (ID 497) is:
 workload_of_jobs P (arrivals_between t1 t2) <=
 service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl

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


}

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

1 subgoal (ID 497)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = false
  ============================
  workload_of_jobs P (arrivals_between t1 t2) <=
  service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl

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


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

1 subgoal (ID 497)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = false
  ============================
  workload_of_jobs P (arrivals_between t1 t2) <=
  service_of_jobs sched P (arrivals_between t1 t2) t1 t_compl

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


unfold workload_of_jobs, service_of_jobs.

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

1 subgoal (ID 507)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = false
  ============================
  \sum_(j <- arrivals_between t1 t2 | P j) job_cost j <=
  \sum_(j <- arrivals_between t1 t2 | P j) service_during sched j t1 t_compl

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


        rewrite big_seq_cond [X in _ X]big_seq_cond.

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

1 subgoal (ID 530)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = false
  ============================
  \sum_(i <- arrivals_between t1 t2 | (i \in arrivals_between t1 t2) && P i)
     job_cost i <=
  \sum_(i <- arrivals_between t1 t2 | (i \in arrivals_between t1 t2) && P i)
     service_during sched i t1 t_compl

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


        rewrite leq_sum //.

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

1 subgoal (ID 539)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = false
  ============================
  forall i : Job,
  (i \in arrivals_between t1 t2) && P i ->
  job_cost i <= service_during sched i t1 t_compl

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


        movej /andP [ARR Pj].

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

1 subgoal (ID 604)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  COMPL : all_jobs_completed_by t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = false
  j : Job
  ARR : j \in arrivals_between t1 t2
  Pj : P j
  ============================
  job_cost j <= service_during sched j t1 t_compl

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


        specialize (COMPL _ ARR Pj).

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

1 subgoal (ID 608)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  j : Job
  COMPL : completed_by j t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = false
  ARR : j \in arrivals_between t1 t2
  Pj : P j
  ============================
  job_cost j <= service_during sched j t1 t_compl

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


        rewrite -[service_during _ _ _ _ ]add0n.

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

1 subgoal (ID 624)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  j : Job
  COMPL : completed_by j t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = false
  ARR : j \in arrivals_between t1 t2
  Pj : P j
  ============================
  job_cost j <= 0 + service_during sched j t1 t_compl

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


        rewrite -(F j t1); try done.

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

1 subgoal (ID 628)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  j : Job
  COMPL : completed_by j t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = false
  ARR : j \in arrivals_between t1 t2
  Pj : P j
  ============================
  job_cost j <=
  service_during sched j 0 t1 + service_during sched j t1 t_compl

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


        rewrite -(big_cat_nat) //=; last first.

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

1 subgoal (ID 668)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  j : Job
  COMPL : completed_by j t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  EQ : (t_compl <= t1) = false
  ARR : j \in arrivals_between t1 t2
  Pj : P j
  ============================
  t1 <= t_compl

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


        move: EQ =>/negP /negP; rewrite -ltnNge; moveEQ.

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

1 subgoal (ID 771)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  j : Job
  COMPL : completed_by j t_compl
  F : forall (j : Job) (t : nat),
      t <= t1 ->
      j \in arrivals_between t1 t2 -> service_during sched j 0 t = 0
  ARR : j \in arrivals_between t1 t2
  Pj : P j
  EQ : t1 < t_compl
  ============================
  t1 <= t_compl

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


          by apply ltnW.

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

No more subgoals.

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


      }

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

No more subgoals.

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


    Qed.

Using the lemmas above, we prove equivalence.
    Corollary all_jobs_have_completed_equiv_workload_eq_service:
      all_jobs_completed_by t_compl
      workload_of_jobs P jobs =
      service_of_jobs sched P jobs t1 t_compl.

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

1 subgoal (ID 73)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  ============================
  all_jobs_completed_by t_compl <->
  workload_of_jobs P jobs = service_of_jobs sched P jobs t1 t_compl

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


    Proof.
      split.

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

2 subgoals (ID 75)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  ============================
  all_jobs_completed_by t_compl ->
  workload_of_jobs P jobs = service_of_jobs sched P jobs t1 t_compl

subgoal 2 (ID 76) is:
 workload_of_jobs P jobs = service_of_jobs sched P jobs t1 t_compl ->
 all_jobs_completed_by t_compl

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


      - by apply all_jobs_have_completed_impl_workload_eq_service.

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

1 subgoal (ID 76)
  
  Task : TaskType
  Job : JobType
  H : JobTask Job Task
  H0 : JobArrival Job
  H1 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  P : Job -> bool
  arrivals_between := arrival_sequence.arrivals_between arr_seq
   : instant -> instant -> seq Job
  completed_by := service.completed_by sched : Job -> instant -> bool
  t1, t2 : instant
  jobs := arrivals_between t1 t2 : seq Job
  t_compl : instant
  all_jobs_completed_by := fun t_compl : instant =>
                           forall j : Job,
                           j \in jobs -> P j -> completed_by j t_compl
   : instant -> Prop
  ============================
  workload_of_jobs P jobs = service_of_jobs sched P jobs t1 t_compl ->
  all_jobs_completed_by t_compl

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


      - by apply workload_eq_service_impl_all_jobs_have_completed.

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

No more subgoals.

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


    Qed.

  End WorkloadServiceAndCompletion.

End IdealModelLemmas.