Library prosa.analysis.facts.busy_interval.busy_interval

Throughout this file, we assume ideal uni-processor schedules.

Existence of Busy Interval for JLFP-models

In this module we derive a sufficient condition for existence of busy intervals for uni-processor for JLFP schedulers.
Consider any type of tasks ...
  Context {Task : TaskType}.
  Context `{TaskCost Task}.

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

Consider any arrival sequence with consistent arrivals.
Next, consider any ideal uni-processor schedule of this arrival sequence ...
... where jobs do not execute before their arrival or after completion.
Assume a given JLFP policy.
  Context `{JLFP_policy Job}.

Further, allow for any work-bearing notion of job readiness.
For simplicity, let's define some local names.
Consider an arbitrary task tsk.
  Variable tsk : Task.

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

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

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

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

  End BasicLemma.

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

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

Then, we prove that there is a job pending at time t2 that has higher or equal priority (with respect to tsk).
    Lemma not_quiet_implies_exists_pending_job:
       j_hp,
        arrives_in arr_seq j_hp
        arrived_between j_hp t1 t2
        hep_job j_hp j
        ¬ job_completed_by j_hp t2.

  End ExistsPendingJob.

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

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

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

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

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

  End ProcessorAlwaysBusy.

In section we prove a few auxiliary lemmas about quiet time and service.
Assume that the schedule is work-conserving ...
... and there are no duplicate job arrivals.
Let t1 be a quiet time.
    Variable t1 : instant.
    Hypothesis H_quiet_time : quiet_time t1.

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

For clarity, we introduce a notion of the total service of jobs released in time interval [t_beg, t_end) during the time interval [t1, t1 + Δ).
We prove that jobs with higher-than-or-equal priority that released before time instant t1 receive no service after time instant t1.
Next we prove that the total service within a "non-quiet" time interval [t1, t1 + Δ) is exactly Δ.
In this section, we show that the length of any busy interval is bounded, as long as there is enough supply to accommodate the workload of tasks with higher or equal priority.
  Section BoundingBusyInterval.

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

... and the priority relation is reflexive and transitive.
Next, we recall the notion of workload of all jobs released in a given interval [t1, t2) that have higher-or-equal priority w.r.t. the job j being analyzed.
With regard to the jobs with higher-or-equal priority that are released in a given interval [t1, t2), we also recall the service received by these jobs in the same interval [t1, t2).
Now we begin the proof. First, we show that the busy interval is bounded.
    Section BoundingBusyInterval.

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

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

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

      End LowerBound.

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

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

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

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

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

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

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

Moreover, the fact that the interval is not quiet also implies that there's more workload requested than service received.
          Lemma busy_interval_too_much_workload:
            hp_workload t1 (t1 + delta) > hp_service t1 (t1 + delta).

Using the two lemmas above, we infer that the workload is larger than the interval length. However, this contradicts the assumption H_workload_is_bounded.
          Corollary busy_interval_workload_larger_than_interval:
            priority_inversion_bound + hp_workload t1 (t1 + delta) > delta.

        End CannotBeBusyForSoLong.

Since the interval cannot remain busy for so long, we prove that the busy interval finishes at some point t2 t1 + delta.
        Lemma busy_interval_is_bounded:
           t2,
            t2 t1 + delta
            busy_interval t1 t2.

      End UpperBound.

    End BoundingBusyInterval.

In this section, we show that from a workload bound we can infer the existence of a busy interval.
Let priority_inversion_bound be a constant that bounds the length of a priority inversion.
Assume that for some positive delta, the sum of requested workload at time t1 + delta and priority inversion is bounded by delta (i.e., the supply).
      Variable delta: duration.
      Hypothesis H_delta_positive: delta > 0.
      Hypothesis H_workload_is_bounded:
         t, priority_inversion_bound + hp_workload t (t + delta) delta.

Next, we assume that job j has positive cost, from which we can infer that there is a time in which j is pending.
      Hypothesis H_positive_cost: job_cost j > 0.

Therefore there must exists a busy interval [t1, t2) that contains the arrival time of j.
      Corollary exists_busy_interval:
         t1 t2,
          t1 job_arrival j < t2
          t2 t1 + delta
          busy_interval t1 t2.


    End BusyIntervalFromWorkloadBound.

If we know that the workload is bounded, we can also use the busy interval to infer a response-time bound.
Let priority_inversion_bound be a constant that bounds the length of a priority inversion.
Assume that for some positive delta, the sum of requested workload at time t1 + delta and priority inversion is bounded by delta (i.e., the supply).
      Variable delta: duration.
      Hypothesis H_delta_positive: delta > 0.
      Hypothesis H_workload_is_bounded:
         t, priority_inversion_bound + hp_workload t (t + delta) delta.

Then, job j must complete by job_arrival j + delta.