Library prosa.analysis.facts.busy_interval.priority_inversion

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

Throughout the file we assume for the classic Liu & Layland model of readiness without jitter and no self-suspensions, where pending jobs are always ready.
Require Import prosa.model.readiness.basic.

In preparation of the derivation of the priority inversion bound, we establish two basic facts on preemption times.
Section PreemptionTimes.

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

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

In addition, we assume the existence of a function mapping a task to its maximal non-preemptive segment ...
... and the existence of a function mapping a job and its progress to a boolean value saying whether this job is preemptable at its current point of execution.
  Context `{JobPreemptable Job}.

Consider any arrival sequence with consistent arrivals.
Next, consider any ideal uniprocessor schedule of this arrival sequence ...
Consider a valid model with bounded nonpreemptive segments.
Then, we show that time 0 is a preemption time.
Also, we show that the first instant of execution is a preemption time.
  Lemma first_moment_is_pt:
     j prt,
      arrives_in arr_seq j
      ~~ scheduled_at sched j prt
      scheduled_at sched j prt.+1
      preemption_time sched prt.+1.

End PreemptionTimes.

Priority inversion is bounded

In this module we prove that any priority inversion that occurs in the model with bounded nonpreemptive segments defined in module prosa.model.schedule.uni.limited.platform.definitions is bounded.
Consider any type of tasks ...
  Context {Task : TaskType}.
  Context `{TaskCost Task}.

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

Consider any arrival sequence with consistent arrivals ...
... and any ideal uniprocessor schedule of this arrival sequence ...
... where jobs do not execute before their arrival or after completion.
Consider a JLFP policy that indicates a higher-or-equal priority relation, and assume that the relation is reflexive and transitive.
In addition, we assume the existence of a function mapping a task to its maximal non-preemptive segment ...
... and the existence of a function mapping a job and its progress to a boolean value saying whether this job is preemptable at its current point of execution.
  Context `{JobPreemptable Job}.

And assume that they define a valid preemption model with bounded nonpreemptive segments.
Next, we assume that the schedule is a work-conserving schedule...
... and the schedule respects the policy defined by the predicate job_preemptable (i.e., bounded nonpreemptive segments).
Let's define some local names for clarity.
Finally, we introduce the notion of the maximal length of (potential) priority inversion at a time instant t, which is defined as the maximum length of nonpreemptive segments among all jobs that arrived so far. Note that the value job_max_nonpreemptive_segment j_lp is at least ε for any job j_lp, so the maximal length of priority inversion cannot be negative.
Next we prove that a priority inversion of a job is bounded by function max_length_of_priority_inversion.
Note that any bound on function max_length_of_priority_inversion will also be a bound on the maximal priority inversion. This bound may be different for different scheduler and/or task models. Thus, we don't define such a bound in this module.
Consider any job j of tsk with positive job cost.
  Variable j : Job.
  Hypothesis H_j_arrives : arrives_in arr_seq j.
  Hypothesis H_job_cost_positive : job_cost_positive j.

Consider any busy interval prefix [t1, t2) of job j.
  Variable t1 t2 : instant.
  Hypothesis H_busy_interval_prefix:
    busy_interval_prefix arr_seq sched j t1 t2.

Processor Executes HEP jobs after Preemption Point

In this section, we prove that at any time instant after any preemption point (inside the busy interval), the processor is always busy scheduling a job with higher or equal priority.
First, we show that the processor at any preemptive point is always busy scheduling a job with higher or equal priority.
Consider an arbitrary preemption time t ∈ [t1,t2).
      Variable t : instant.
      Hypothesis H_t_in_busy_interval : t1 t < t2.
      Hypothesis H_t_preemption_time : preemption_time sched t.

First note since t lies inside the busy interval, the processor cannot be idle at time t.
      Lemma instant_t_is_not_idle:
        ¬ is_idle sched t.

Next we consider two cases: (1) when t is less than t2 - 1 and (2) t is equal to t2 - 1.
      Lemma t_lt_t2_or_t_eq_t2:
        t < t2.-1 t = t2.-1.

In case when t < t2 - 1, we use the fact that time instant t+1 is not a quiet time. The later implies that there is a pending higher-or-equal priority job at time t. Thus, the assumption that the schedule respects priority policy at preemption points implies that the scheduled job must have a higher-or-equal priority.
      Lemma scheduled_at_preemption_time_implies_higher_or_equal_priority_lt:
        t < t2.-1
         jhp,
          scheduled_at sched jhp t
          hep_job jhp j.

In case when t = t2 - 1, we cannot use the same proof since t+1 = t2, but t2 is a quiet time. So we do a similar case analysis on the fact that t1 = t t1 < t.
      Lemma scheduled_at_preemption_time_implies_higher_or_equal_priority_eq:
        t = t2.-1
         jhp,
          scheduled_at sched jhp t
          hep_job jhp j.

By combining the above facts we conclude that a job that is scheduled at time t has higher-or-equal priority.
      Corollary scheduled_at_preemption_time_implies_higher_or_equal_priority:
         jhp,
          scheduled_at sched jhp t
          hep_job jhp j.

Since a job that is scheduled at time t has higher-or-equal priority, by properties of a busy interval it cannot arrive before time instant t1.
      Lemma scheduled_at_preemption_time_implies_arrived_between_within_busy_interval:
         jhp,
          scheduled_at sched jhp t
          arrived_between jhp t1 t2.

From the above lemmas we prove that there is a job jhp that is (1) scheduled at time t, (2) has higher-or-equal priority, and (3) arrived between t1 and t2.
      Corollary not_quiet_implies_exists_scheduled_hp_job_at_preemption_point:
         j_hp,
          arrived_between j_hp t1 t2
          hep_job j_hp j
          job_scheduled_at j_hp t.

    End ProcessorBusyWithHEPJobAtPreemptionPoints.

Next we prove that every nonpreemptive segment always begins with a preemption time ...
    Lemma scheduling_of_any_segment_starts_with_preemption_time:
       j t,
        job_scheduled_at j t
         pt,
          job_arrival j pt t
          preemption_time sched pt
          ( t', pt t' t job_scheduled_at j t').

... and the fact that at any time instant after a preemption point the processor is always busy with a job with higher or equal priority.
    Lemma not_quiet_implies_exists_scheduled_hp_job_after_preemption_point:
       tp t,
        preemption_time sched tp
        t1 tp < t2
        tp t < t2
         j_hp,
          arrived_between j_hp t1 t.+1
          hep_job j_hp j
          job_scheduled_at j_hp t.

Now, suppose there exists some constant K that bounds the distance to a preemption time from the beginning of the busy interval.
    Variable K : duration.
    Hypothesis H_preemption_time_exists:
       pr_t, preemption_time sched pr_t t1 pr_t t1 + K.

Then we prove that the processor is always busy with a job with higher-or-equal priority after time instant t1 + K.
    Lemma not_quiet_implies_exists_scheduled_hp_job:
       t,
        t1 + K t < t2
         j_hp,
          arrived_between j_hp t1 t.+1
          hep_job j_hp j
          job_scheduled_at j_hp t.

  End PreemptionTimeAndPriorityInversion.

Preemption Time Exists

In this section we prove that the function max_length_of_priority_inversion indeed upper bounds the priority inversion length.
  Section PreemptionTimeExists.

First we prove that if a job with higher-or-equal priority is scheduled at a quiet time t+1 then this is the first time when this job is scheduled.
    Lemma hp_job_not_scheduled_before_quiet_time:
       jhp t,
        quiet_time arr_seq sched j t.+1
        job_scheduled_at jhp t.+1
        hep_job jhp j
        ~~ job_scheduled_at jhp t.

Also, we show that lower-priority jobs that are scheduled inside the busy-interval prefix [t1,t2) must arrive before that interval.
    Lemma low_priority_job_arrives_before_busy_interval_prefix:
       jlp t,
        t1 t < t2
        job_scheduled_at jlp t
        ~~ hep_job jlp j
        job_arrival jlp < t1.

Moreover, we show that lower-priority jobs that are scheduled inside the busy-interval prefix [t1,t2) must be scheduled before that interval.
    Lemma low_priority_job_scheduled_before_busy_interval_prefix:
       jlp t,
        t1 t < t2
        job_scheduled_at jlp t
        ~~ hep_job jlp j
         t', t' < t1 job_scheduled_at jlp t'.

Thus, there must be a preemption time in the interval t1, t1 + max_priority_inversion t1. That is, if a job with higher-or-equal priority is scheduled at time instant t1, then t1 is a preemption time. Otherwise, if a job with lower priority is scheduled at time t1, then this jobs also should be scheduled before the beginning of the busy interval. So, the next preemption time will be no more than max_priority_inversion t1 time units later.
We proceed by doing a case analysis.
    Section CaseAnalysis.

(1) Case when the schedule is idle at time t1.
      Section Case1.

Assume that the schedule is idle at time t1.
        Hypothesis H_is_idle : is_idle sched t1.

Then time instant t1 is a preemption time.
        Lemma preemption_time_exists_case1:
           pr_t,
            preemption_time sched pr_t
            t1 pr_t t1 + max_length_of_priority_inversion j t1.

      End Case1.

(2) Case when a job with higher-or-equal priority is scheduled at time t1.
      Section Case2.

Assume that a job jhp with higher-or-equal priority is scheduled at time t1.
        Variable jhp : Job.
        Hypothesis H_jhp_is_scheduled : scheduled_at sched jhp t1.
        Hypothesis H_jhp_hep_priority : hep_job jhp j.

Then time instant t1 is a preemption time.
        Lemma preemption_time_exists_case2:
           pr_t,
            preemption_time sched pr_t
            t1 pr_t t1 + max_length_of_priority_inversion j t1.

      End Case2.

(3) Case when a job with lower priority is scheduled at time t1.
      Section Case3.

Assume that a job jhp with lower priority is scheduled at time t1.
        Variable jlp : Job.
        Hypothesis H_jlp_is_scheduled : scheduled_at sched jlp t1.
        Hypothesis H_jlp_low_priority : ~~ hep_job jlp j.

To prove the lemma in this case we need a few auxiliary facts about the first preemption point of job jlp.
        Section FirstPreemptionPointOfjlp.

Let's denote the progress of job jlp at time t1 as progr_t1.
          Let progr_t1 := service sched jlp t1.

Consider the first preemption point of job jlp after progr_t1.
          Variable fpt : instant.
          Hypothesis H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt).
          Hypothesis H_fpt_is_first_preemption_point:
             ρ,
              progr_t1 ρ progr_t1 + (job_max_nonpreemptive_segment jlp - ε)
              job_preemptable jlp ρ
              service sched jlp t1 + fpt ρ.

For convenience we also assume the following inequality holds.
          Hypothesis H_progr_le_max_nonp_segment:
            progr_t1 progr_t1 + fpt progr_t1 + (job_max_nonpreemptive_segment jlp - ε).

First we show that fpt is indeed the first preemption point after progr_t1.
          Lemma no_intermediate_preemption_point:
             ρ,
              progr_t1 ρ < progr_t1 + fpt
              ~~ job_preemptable jlp ρ.

Thanks to the fact that the scheduler respects the notion of preemption points we show that jlp is continuously scheduled in time interval [t1, t1 + fpt).
          Lemma continuously_scheduled_between_preemption_points:
             t',
              t1 t' < t1 + fpt
              job_scheduled_at jlp t'.
Thus, job jlp reaches its preemption point at time instant t1 + fpt, which implies that time instant t1 + fpt is a preemption time.
          Lemma first_preemption_time:
            preemption_time sched (t1 + fpt).

          Lemma preemption_time_le_max_len_of_priority_inversion:
            t1 t1 + fpt t1 + max_length_of_priority_inversion j t1.

        End FirstPreemptionPointOfjlp.

Next we combine the facts above to conclude the lemma.
        Lemma preemption_time_exists_case3:
           pr_t,
            preemption_time sched pr_t
            t1 pr_t t1 + max_length_of_priority_inversion j t1.

      End Case3.

    End CaseAnalysis.

By doing the case analysis, we show that indeed there is a preemption time in time interval [t1, t1 + max_length_of_priority_inversion j t1].