Library prosa.analysis.facts.transform.swaps

In this file, we establish invariants about schedules in which two allocations have been swapped, as for instance it is done in the classic EDF optimality proof.

Section SwappedFacts.
For any given type of jobs...
  Context {Job : JobType}.
... any given type of processor states:
  Context {PState: Type}.
  Context `{ProcessorState Job PState}.

...consider any given reference schedule.
  Variable sched: schedule PState.

Suppose we are given two specific times t1 and t2.
  Variable t1 t2: instant.

In the following, let sched' denote the schedule in which the allocations at t1 and t2 have been swapped.
  Let sched' := swapped sched t1 t2.

First, we note that the trivial case where t1 == t2 is not interesting because then the two schedules are identical.
  Lemma trivial_swap:
    t1 = t2
     t,
      sched t = sched' t.

In this trivial case, the amount of service received hence is obviously always identical.
  Lemma trivial_swap_service_invariant:
    t1 = t2
     t j,
      service sched j t = service sched' j t.

In any case, the two schedules do not differ at non-swapped times.
  Lemma swap_other_times_invariant:
     t,
      t t1
      t t2
      sched t = sched' t.

By definition, if a job is scheduled at t2 in the original schedule, then it is found at t1 in the new schedule.
  Lemma swap_job_scheduled_t1:
     j,
      scheduled_at sched' j t1 =
      scheduled_at sched j t2.

Similarly, a job scheduled at t1 in the original schedule is scheduled at t2 after the swap.
  Lemma swap_job_scheduled_t2:
     j,
      scheduled_at sched' j t2 =
      scheduled_at sched j t1.

If a job is scheduled at any time not involved in the swap, then it remains scheduled at that time in the new schedule.
  Lemma swap_job_scheduled_other_times:
     j t,
      t1 != t
      t2 != t
      scheduled_at sched' j t =
      scheduled_at sched j t.

To make case analysis more convenient, we summarize the preceding three lemmas as a disjunction.
  Corollary swap_job_scheduled_cases:
     j t,
      scheduled_at sched' j t
      scheduled_at sched' j t = scheduled_at sched j t
      
      t = t1 scheduled_at sched' j t = scheduled_at sched j t2
      
      t = t2 scheduled_at sched' j t = scheduled_at sched j t1.

From this, we can easily conclude that no jobs have appeared out of thin air: if a job scheduled at some time in the new schedule, then it was also scheduled at some time in the original schedule.
  Corollary swap_job_scheduled:
     j t,
      scheduled_at sched' j t
       t',
        scheduled_at sched j t'.

Mirroring swap_job_scheduled_cases above, we also state a disjunction for case analysis under the premise that a job is scheduled in the original schedule.
  Lemma swap_job_scheduled_original_cases:
     j t,
      scheduled_at sched j t
      scheduled_at sched' j t = scheduled_at sched j t
      
      t = t1 scheduled_at sched' j t2 = scheduled_at sched j t
      
      t = t2 scheduled_at sched' j t1 = scheduled_at sched j t.

Thus, we can conclude that no jobs are lost: if a job is scheduled at some point in the original schedule, then it is also scheduled at some point in the new schedule.
  Corollary swap_job_scheduled_original:
     j t,
      scheduled_at sched j t
       t',
        scheduled_at sched' j t'.

In the following, we lift the above basic invariants to statements about invariants about the cumulative amount of service received.
To avoid dealing with symmetric cases, assume in the following that t1 and t2 are ordered.
  Hypothesis H_well_ordered: t1 t2.

As another trivial invariant, we observe that nothing has changed before time t1.
  Lemma swap_before_invariant:
     t,
      t < t1
      sched t = sched' t.

Similarly, nothing has changed after time t2.
  Lemma swap_after_invariant:
     t,
      t2 < t
      sched t = sched' t.

Thus, we observe that, before t1, the two schedules are identical with regard to the service received by any job because they are identical.
  Corollary service_before_swap_invariant:
     t,
      t t1
       (j : Job),
        service sched j t = service sched' j t.

Likewise, we observe that, *after* t2, the swapped schedule again does not differ with regard to the service received by any job.
  Lemma service_after_swap_invariant:
     t,
      t2 < t
       (j : Job),
        service sched j t = service sched' j t.

Finally, we note that, trivially, jobs that are not involved in the swap receive invariant service.
  Lemma service_of_others_invariant:
     t j,
      ~~ scheduled_in j (sched t1)
      ~~ scheduled_in j (sched t2)
      service sched j t = service sched' j t.

End SwappedFacts.

In the following section, we observe that, if the original schedule satisfies certain properties, then so does the new scheduled obtained by swapping the allocations at times t1 and t2.
For any given type of jobs...
  Context {Job : JobType} `{JobCost Job} `{JobDeadline Job} `{JobArrival Job}.
... any given type of processor states:
  Context {PState: eqType}.
  Context `{ProcessorState Job PState}.

...consider any given reference schedule.
  Variable sched: schedule PState.

Suppose we are given two specific times t1 and t2...
  Variable t1 t2: instant.

...such that t1 is no later than t2.
  Hypothesis H_order: t1 t2.

We let sched' denote the schedule in which the allocations at t1 and t2 have been swapped.
  Let sched' := swapped sched t1 t2.

First, we observe that if jobs never accumulate more service than required, then that's still the case after the swap.
  Lemma swapped_service_bound:
    ( j t, service sched j t job_cost j)
    ( j t, service sched' j t job_cost j).

From the above service bound, we conclude that, if if completed jobs don't execute in the original schedule, then that's still the case after the swap, assuming an ideal unit-service model (i.e., scheduled jobs receive exactly one unit of service).
Suppose all jobs in the original schedule come from some arrival sequence...
...then that's still the case in the new schedule.
In the next section, we consider a special case of swaps: namely, when the job moved earlier has an earlier deadline than the job being moved to a later allocation, assuming that the former has not yet missed its deadline, which is the core transformation of the classic EDF optimality proof.
Section EDFSwap.
For any given type of jobs with costs and deadlines...
  Context {Job : JobType} `{JobCost Job} `{JobDeadline Job}.
... any given type of processor states...
  Context {PState: eqType}.
  Context `{ProcessorState Job PState}.

...consider a given reference schedule...
  Variable sched: schedule PState.

...in which complete jobs don't execute...
...and scheduled jobs always receive service.
Suppose we are given two specific times t1 and t2...
  Variable t1 t2: instant.

...that are ordered.
  Hypothesis H_well_ordered: t1 t2.

Further, assume that, if there are jobs scheduled at times t1 and t2, then they either have the same deadline or violate EDF, ...
  Hypothesis H_not_EDF:
     j1 j2,
      scheduled_at sched j1 t1
      scheduled_at sched j2 t2
      job_deadline j1 job_deadline j2.

...and that we don't move idle times or deadline misses earlier, i.e., if t1 is not an idle time, then neither is t2 and whatever job is scheduled at time t2 has not yet missed its deadline.
  Hypothesis H_no_idle_time_at_t2:
     j1,
      scheduled_at sched j1 t1
       j2, scheduled_at sched j2 t2 job_deadline j2 > t2.

Consider the schedule obtained from swapping the allocations at times t1 and t2.
  Let sched' := swapped sched t1 t2.

The key property of this transformation is that, for any job that meets its deadline in the original schedule, we have not introduced any deadline misses, which we establish by considering a number of different cases.
Consider any job...
    Variable j: Job.

... that meets its deadline in the original schedule.
    Hypothesis H_deadline_met: job_meets_deadline sched j.

First we observe that jobs that are not involved in the swap still meet their deadlines.
Next, we observe that a swap is unproblematic for the job scheduled at time t2.
Finally, we observe is also unproblematic for the job that is moved to a later allocation.
From the above case analysis, we conclude that no deadline misses are introduced in the schedule obtained from swapping the allocations at times t1 and t2.