Library prosa.analysis.facts.model.ideal.service_of_jobs

Service Received by Sets of Jobs in Uniprocessor Ideal-Progress Schedules

In this file, we establish a fact about the service received by sets of jobs in the presence of idle times on a uniprocessor under the ideal-progress assumption (i.e., that a scheduled job necessarily receives nonzero service).
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.
Allow for any uniprocessor model that ensures ideal progress.
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 : pred Job.

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 arr_seq 0 t2) t1 t2 < t2 - t1
       t, t1 t < t2 is_idle arr_seq sched t.
  Proof.
    movet1 t2 SERV.
    destruct (t1 t2) eqn:LE; last first.
    { move: LE ⇒ /negP/negP; rewrite -ltnNge.
      moveLT; apply ltnW in LT; rewrite -subn_eq0 in LT.
      by move: LT ⇒ /eqPin SERV; rewrite ltn0 in SERV.
    }
    have EX: δ, t2 = t1 + δ.
    { by (t2 - t1); rewrite subnKC // ltnW. }
    move: EX ⇒ [δ EQ]; subst t2; clear LE.
    rewrite {3}[t1 + δ]addnC -addnBA // subnn addn0 in SERV.
    rewrite /service_of_jobs exchange_big //= in SERV.
    apply sum_le_summation_range in SERV.
    move: SERV ⇒ [x [/andP [GEx LEx] L]].
     x; split; first by apply/andP; split.
    apply/negPn/negP; rewrite is_nonidle_iff //= ⇒ [[j SCHED]].
    move: L ⇒ /eqP; rewrite sum_nat_eq0_nat filter_predT ⇒ /allP L.
    have /eqP: service_at sched j x == 0.
    { apply/L/arrived_between_implies_in_arrivals ⇒ //.
      rewrite /arrived_between; apply/andP; split ⇒ //.
      have: job_arrival j x by apply: H_jobs_must_arrive_to_execute.
      lia. }
    by rewrite -no_service_not_scheduled // ⇒ /negP.
  Qed.

End IdealModelLemmas.