Library rt.model.jitter.interference

Require Import rt.util.all.
Require Import rt.model.jitter.task rt.model.jitter.job rt.model.jitter.schedule
               rt.model.jitter.priority rt.model.jitter.workload.

Module Interference.

  Import ScheduleOfSporadicTaskWithJitter Priority Workload.

  (* We import some of the basic definitions, but we need to re-define almost everything
     since the definition of backlogged (and thus the definition of interference)
     changes with jitter. *)

  Require Import rt.model.basic.interference.
  Export Interference.

  Section InterferenceDefs.

    Context {sporadic_task: eqType}.
    Context {Job: eqType}.
    Variable job_cost: Job time.
    Variable job_task: Job sporadic_task.
    Variable job_jitter: Job time.

    (* Assume any job arrival sequence...*)
    Context {arr_seq: arrival_sequence Job}.

    (* ... and any schedule. *)
    Context {num_cpus: nat}.
    Variable sched: schedule num_cpus arr_seq.

    (* Consider any job j that incurs interference. *)
    Variable j: JobIn arr_seq.

    (* Recall the definition of backlogged (pending and not scheduled). *)
    Let job_is_backlogged (t: time) :=
      backlogged job_cost job_jitter sched j t.

    Section TotalInterference.

      (* First, we define the total interference incurred by job j during [t1, t2) as the cumulative time in which
         j is backlogged in this interval. *)

      Definition total_interference (t1 t2: time) :=
        \sum_(t1 t < t2) job_is_backlogged t.

    End TotalInterference.

    Section JobInterference.

      (* Let job_other be a job that interferes with j. *)
      Variable job_other: JobIn arr_seq.

      (* We define the total interference caused by job_other during [t1, t2) as the cumulative service
         received by job_other while j is backlogged. *)

      Definition job_interference (t1 t2: time) :=
        \sum_(t1 t < t2)
          \sum_(cpu < num_cpus)
            (job_is_backlogged t &&
            scheduled_on sched job_other cpu t).

    End JobInterference.

    Section TaskInterference.

      (* In order to define task interference, consider any interfering task tsk_other. *)
      Variable tsk_other: sporadic_task.

      (* We define the total interference caused by tsk during [t1, t2) as the cumulative service
         received by tsk while j is backlogged. *)

      Definition task_interference (t1 t2: time) :=
        \sum_(t1 t < t2)
          \sum_(cpu < num_cpus)
            (job_is_backlogged t &&
            task_scheduled_on job_task sched tsk_other cpu t).

    End TaskInterference.

    Section TaskInterferenceJobList.

      Variable tsk_other: sporadic_task.

      Definition task_interference_joblist (t1 t2: time) :=
        \sum_(j <- jobs_scheduled_between sched t1 t2 | job_task j == tsk_other)
         job_interference j t1 t2.

    End TaskInterferenceJobList.

    Section BasicLemmas.

      (* Interference cannot be larger than the considered time window. *)
      Lemma total_interference_le_delta :
         t1 t2,
          total_interference t1 t2 t2 - t1.

      Lemma job_interference_le_service :
         j_other t1 t2,
          job_interference j_other t1 t2 service_during sched j_other t1 t2.

      Lemma task_interference_le_workload :
         tsk t1 t2,
          task_interference tsk t1 t2 workload job_task sched tsk t1 t2.

    End BasicLemmas.

    Section InterferenceSequentialJobs.

      (* If jobs are sequential, ... *)
      Hypothesis H_sequential_jobs: sequential_jobs sched.

      (* ... then the interference incurred by a job in an interval
         of length delta is at most delta. *)

      Lemma job_interference_le_delta :
         j_other t1 delta,
          job_interference j_other t1 (t1 + delta) delta.
    End InterferenceSequentialJobs.

    (* The sequential per-job interference bounds the actual interference. *)
    Section BoundUsingPerJobInterference.

      Lemma interference_le_interference_joblist :
         tsk t1 t2,
          task_interference tsk t1 t2 task_interference_joblist tsk t1 t2.

    End BoundUsingPerJobInterference.

  End InterferenceDefs.

End Interference.