Library rt.analysis.apa.bertogna_edf_theory

Require Import rt.util.all rt.util.divround.
Require Import rt.model.apa.task rt.model.apa.job rt.model.apa.task_arrival
               rt.model.apa.schedule rt.model.apa.platform rt.model.apa.interference
               rt.model.apa.workload rt.model.apa.schedulability rt.model.apa.priority
               rt.model.apa.platform rt.model.apa.response_time
               rt.model.apa.affinity rt.model.apa.constrained_deadlines.
Require Import rt.analysis.apa.workload_bound rt.analysis.apa.interference_bound_edf.

Module ResponseTimeAnalysisEDF.

  Export Job SporadicTaskset ScheduleOfSporadicTask Workload Schedulability ResponseTime
         Priority SporadicTaskArrival WorkloadBound InterferenceBoundEDF
         Interference Platform Affinity ConstrainedDeadlines.

  (* In this section, we prove that any fixed point in the APA-reduction of
     Bertogna and Cirinei's RTA for EDF scheduling is a safe response-time bound.
     This result corresponds to Lemma 10 in the revised version of the APA paper:
     http://www.mpi-sws.org/~bbb/papers/pdf/ecrts13a-rev1.pdf *)

  Section ResponseTimeBound.

    Context {sporadic_task: eqType}.
    Variable task_cost: sporadic_task time.
    Variable task_period: sporadic_task time.
    Variable task_deadline: sporadic_task time.

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

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

    (* ... in which jobs arrive sporadically and have valid parameters. *)
    Hypothesis H_sporadic_tasks:
      sporadic_task_model task_period arr_seq job_task.
    Hypothesis H_valid_job_parameters:
       (j: JobIn arr_seq),
        valid_sporadic_job task_cost task_deadline job_cost job_deadline job_task j.

    (* Consider a task set ts where all tasks have valid parameters
       and constrained deadlines, ... *)

    Variable ts: taskset_of sporadic_task.
    Hypothesis H_valid_task_parameters:
      valid_sporadic_taskset task_cost task_period task_deadline ts.
    Hypothesis H_constrained_deadlines:
       tsk, tsk ts task_deadline tsk task_period tsk.

    (* ... and assume that all jobs in the arrival sequence come from the task set. *)
    Hypothesis H_all_jobs_from_taskset:
       (j: JobIn arr_seq), job_task j ts.

    (* Also assume that every task has a non-empty processor affinity alpha. *)
    Context {num_cpus: nat}.
    Variable alpha: task_affinity sporadic_task num_cpus.

    (* Next, consider any schedule such that...*)
    Variable sched: schedule num_cpus arr_seq.

    (* ...jobs are sequential and do not execute before their
       arrival times nor longer than their execution costs. *)

    Hypothesis H_sequential_jobs: sequential_jobs sched.
    Hypothesis H_jobs_must_arrive_to_execute:
      jobs_must_arrive_to_execute sched.
    Hypothesis H_completed_jobs_dont_execute:
      completed_jobs_dont_execute job_cost sched.

    (* Assume that the schedule is an work-conserving APA schedule that
       enforces EDF priorities. *)

    Hypothesis H_respects_affinity: respects_affinity job_task sched alpha.
    Hypothesis H_work_conserving: apa_work_conserving job_cost job_task sched alpha.
    Hypothesis H_edf_policy:
      enforces_JLDP_policy_under_weak_APA job_cost job_task sched alpha (EDF job_deadline).

    (* Let's define some local names to avoid passing many parameters. *)
    Let no_deadline_is_missed_by_tsk (tsk: sporadic_task) :=
      task_misses_no_deadline job_cost job_deadline job_task sched tsk.
    Let response_time_bounded_by (tsk: sporadic_task) :=
      is_response_time_bound_of_task job_cost job_task tsk sched.

    (* Now we consider the response-time recurrence. In the computation of
       the response-time bound, we assume that each task under analysis has
       a non-empty subaffinity alpha'.
       Note that the notation |...| expresses the cardinality of the set. *)

    Variable alpha': task_affinity sporadic_task num_cpus.
    Hypothesis H_affinity_subset: tsk, tsk ts is_subaffinity (alpha' tsk) (alpha tsk).
    Hypothesis H_at_least_one_cpu : tsk, tsk ts #|alpha' tsk| > 0.

    (* Assume that a response-time bound R is known...  *)
    Let task_with_response_time := (sporadic_task × time)%type.
    Variable rt_bounds: seq task_with_response_time.

    (* ...for any task in the task set, ... *)
    Hypothesis H_rt_bounds_contains_all_tasks: unzip1 rt_bounds = ts.

    (* ... where R is a fixed-point of the response-time recurrence (with
           alpha' as the reduced affinity mask), ... *)

    Let I (tsk: sporadic_task) (delta: time) :=
      total_interference_bound_edf task_cost task_period task_deadline alpha
                                   tsk (alpha' tsk) rt_bounds delta.
    Hypothesis H_response_time_is_fixed_point :
       tsk R,
        (tsk, R) rt_bounds
        R = task_cost tsk + div_floor (I tsk R) #|alpha' tsk|.

    (* ..., and R is no larger than the deadline. *)
    Hypothesis H_tasks_miss_no_deadlines:
       tsk R,
        (tsk, R) rt_bounds R task_deadline tsk.

    (* In order to prove that R is a response-time bound, we first provide some lemmas. *)
    Section Lemmas.

      (* Let (tsk, R) be any task to be analyzed, with its response-time bound R. *)
      Variable tsk: sporadic_task.
      Variable R: time.
      Hypothesis H_tsk_R_in_rt_bounds: (tsk, R) rt_bounds.

      (* Consider any job j of tsk ... *)
      Variable j: JobIn arr_seq.
      Hypothesis H_job_of_tsk: job_task j = tsk.

      (* ... that did not complete on time, ... *)
      Hypothesis H_j_not_completed: ¬ completed job_cost sched j (job_arrival j + R).

      (* ... and that is the first job not to satisfy its response-time bound. *)
      Hypothesis H_all_previous_jobs_completed_on_time :
         (j_other: JobIn arr_seq) tsk_other R_other,
          job_task j_other = tsk_other
          (tsk_other, R_other) rt_bounds
          job_arrival j_other + R_other < job_arrival j + R
          completed job_cost sched j_other (job_arrival j_other + R_other).

      (* Let's call x the interference incurred by job j due to tsk_other, ...*)
      Let x (tsk_other: sporadic_task) :=
        task_interference job_cost job_task sched alpha j
                          tsk_other (job_arrival j) (job_arrival j + R).

      (* and X the total interference incurred by job j due to any task. *)
      Let X := total_interference job_cost sched j (job_arrival j) (job_arrival j + R).

      (* Recall Bertogna and Cirinei's workload bound ... *)
      Let workload_bound (tsk_other: sporadic_task) (R_other: time) :=
        W task_cost task_period tsk_other R_other R.

      (*... and the EDF-specific bound, ... *)
      Let edf_specific_bound (tsk_other: sporadic_task) (R_other: time) :=
        edf_specific_interference_bound task_cost task_period task_deadline tsk tsk_other R_other.

      (* ... which combined form the interference bound. *)
      Let interference_bound (tsk_other: sporadic_task) (R_other: time) :=
        interference_bound_edf task_cost task_period task_deadline tsk R (tsk_other, R_other).

      (* Based on the definition of a different task in subaffinity alpha', ... *)
      Let other_task_in alpha' := different_task_in alpha tsk alpha'.

      (* ...let (other_tasks_in alpha') denote the set of tasks that are different from tsk
         and that can be scheduled on subaffinity alpha'. *)

      Let other_tasks_in alpha' :=
        [seq tsk_other <- ts | other_task_in (alpha' tsk) tsk_other].

      (* Now we establish results the interfering tasks. *)
      Section LemmasAboutInterferingTasks.

        (* Let (tsk_other, R_other) be any pair of higher-priority task and
           response-time bound computed in previous iterations. *)

        Variable tsk_other: sporadic_task.
        Variable R_other: time.
        Hypothesis H_response_time_of_tsk_other: (tsk_other, R_other) rt_bounds.

        (* Note that tsk_other is in the task set, ...*)
        Lemma bertogna_edf_tsk_other_in_ts: tsk_other ts.

        (* ... and R_other is larger than the cost of tsk_other. *)
        Lemma bertogna_edf_R_other_ge_cost :
          R_other task_cost tsk_other.

        (* Since tsk_other cannot interfere more than it executes, we show that
           the interference caused by tsk_other is no larger than workload bound W. *)

        Lemma bertogna_edf_workload_bounds_interference :
          x tsk_other workload_bound tsk_other R_other.

        (* Recall that the edf-specific interference bound also holds for tsk_other. *)
        Lemma bertogna_edf_specific_bound_holds :
          x tsk_other edf_specific_bound tsk_other R_other.

      End LemmasAboutInterferingTasks.

      (* Next we prove some lemmas that help to derive a contradiction.*)
      Section DerivingContradiction.

      (* 0) Since job j did not complete by its response time bound, it follows that
            the total interference X exceeds  R - e_k + 1. *)

      Lemma bertogna_edf_too_much_interference : X R - task_cost tsk + 1.

      (* 1) Next, we prove that during the scheduling window of j, any job that is
            scheduled while j is backlogged comes from a different task.
            This follows from the fact that j is the first job not to complete
            by its response-time bound, so previous jobs of j's task must have
            completed by their periods and cannot be pending. *)

      Lemma bertogna_edf_interference_by_different_tasks :
         t j_other,
          job_arrival j t < job_arrival j + R
          backlogged job_cost sched j t
          scheduled sched j_other t
          job_task j_other tsk.

      (* 2) In order to use the lemmas in constrained_deadlines.v, we show that
            all jobs released before the end of the interval complete by their
            periods. This follows trivially from the hypothesis that all jobs
            before (job_arrival j + R) complete by their response-time bounds. 
            With this lemma, we can conclude that during job j's scheduling
            window there cannot be multiple pending jobs of each task.*)

      Lemma bertogna_edf_all_previous_jobs_complete_by_their_period:
         t (j0: JobIn arr_seq),
          t < job_arrival j + R
          job_arrival j0 + task_period (job_task j0) t
          completed job_cost sched j0
             (job_arrival j0 + task_period (job_task j0)).

      (* 3) Next, we prove that the sum of the interference of each task is equal to the
            total interference multiplied by the number of processors in tsk's *actual*
            affinity. This holds because interference only occurs when all processors in
            the affinity are busy.
            With this lemma we can relate per-task interference with the total interference
            incurred by j (backlogged time). *)

      Lemma bertogna_edf_all_cpus_in_affinity_busy :
        \sum_(tsk_k <- other_tasks_in alpha) x tsk_k = X × #|alpha tsk|.

      (* 4) Recall that the reduction-based analysis considers only the interfering tasks
            within subaffinity (alpha' tsk), as opposed to (alpha tsk). Therefore, we must
            reason about the task interference wihin (alpha' tsk).
            We begin by showing that the cumulative interference within (alpha' tsk) is at
            least the total interference multiplied by the number of processors in (alpha' tsk). *)

      Lemma bertogna_edf_all_cpus_in_subaffinity_busy :
        \sum_(tsk_k <- other_tasks_in alpha') x tsk_k X × #|alpha' tsk|.

      (* Let's define a predicate for whether a task is scheduled on (alpha tsk). *)
      Let scheduled_on_alpha_tsk := fun t tsk_k
        task_scheduled_on_affinity job_task sched (alpha tsk) tsk_k t.

      (* 5) Now we prove that, at all times that j is backlogged, the number
            of tasks whose affinity intersects (alpha' tsk) that are in fact
            scheduled on (alpha' tsk) is at least the size of (alpha' tsk).
            This is required to prove lemma (6). *)

      Lemma bertogna_edf_alpha'_is_full:
         t,
          job_arrival j t < job_arrival j + R
          backlogged job_cost sched j t
          count (scheduled_on_alpha_tsk t) (other_tasks_in alpha') #|alpha' tsk|.

      (* Before stating the next lemma, let (num_tasks_exceeding delta) be the
         number of interfering tasks that can execute on (alpha' tsk) whose
         interference x is larger than delta. *)

      Let num_tasks_exceeding delta := count (fun ix i delta) (other_tasks_in alpha').

      (* 6) Now we prove that, for any delta, if (num_task_exceeding delta > 0), then the
            cumulative interference caused by the complementary set of interfering tasks fills
            the remaining, not-completely-full (|alpha' tsk| - num_tasks_exceeding delta) processors. *)

      Lemma bertogna_edf_interference_in_non_full_processors :
         delta,
          0 < num_tasks_exceeding delta < #|alpha' tsk|
          \sum_(i <- other_tasks_in alpha' | x i < delta) x i delta × (#|alpha' tsk| - num_tasks_exceeding delta).

      (* 7) Based on lemma (6), we prove that, for any interval delta, if the sum of per-task
            interference exceeds (delta * |alpha' tsk|), the same applies for the
            sum of the minimum of the interference and delta. *)

      Lemma bertogna_edf_minimum_exceeds_interference :
         delta,
          \sum_(tsk_k <- other_tasks_in alpha') x tsk_k delta × #|alpha' tsk|
             \sum_(tsk_k <- other_tasks_in alpha') minn (x tsk_k) delta
             delta × #|alpha' tsk|.

      (* 8) Note that lemma (7) only refers to interference within subaffinity (alpha' tsk).
            To reason about the interference incurred by job j in its complete affinity,
            we prove that an exceeding interference on affinity (alpha tsk) also
            implies an exceeding interference on the subaffinity (alpha' tsk). *)

      Lemma bertogna_edf_interference_on_subaffinity :
         delta,
          \sum_(tsk_k <- other_tasks_in alpha) x tsk_k delta × #|alpha tsk|
          \sum_(tsk_k <- other_tasks_in alpha') x tsk_k delta × #|alpha' tsk|.

      (* 9) Next, using lemmas (0), (3), (7) and (8), we prove that the reduction-based
            interference bound is not enough to cover the sum of the minima over all tasks
            (artifact of the proof by contradiction). *)

      Lemma bertogna_edf_sum_exceeds_total_interference:
        \sum_((tsk_other, R_other) <- rt_bounds | other_task_in (alpha' tsk) tsk_other)
          minn (x tsk_other) (R - task_cost tsk + 1) > I tsk R.

      (* 10) After concluding that the sum of the minima exceeds (R - e_i + 1),
            we prove that there exists a tuple (tsk_k, R_k) that satisfies
            min (x_k, R - e_i + 1) > min (W_k, I_edf, R - e_i + 1).
            This implies that either x_k > W_k or x_k > I_edf, which is a contradiction,
            since both W_k and I_edf are valid task interference bounds. *)

      Lemma bertogna_edf_exists_task_that_exceeds_bound :
         tsk_other R_other,
          (tsk_other, R_other) rt_bounds
          (minn (x tsk_other) (R - task_cost tsk + 1) > interference_bound tsk_other R_other).

      End DerivingContradiction.

    End Lemmas.

    (* Using the lemmas above, we now prove that any response-time bound
       obtained from the analysis is safe. *)

    Section MainProof.

      (* Let (tsk, R) be any task to be analyzed, with its response-time bound R. *)
      Variable tsk: sporadic_task.
      Variable R: time.
      Hypothesis H_tsk_R_in_rt_bounds: (tsk, R) rt_bounds.

      (* Then, R bounds the response-time of tsk. *)
      Theorem bertogna_cirinei_response_time_bound_edf :
        response_time_bounded_by tsk R.

    End MainProof.

  End ResponseTimeBound.

End ResponseTimeAnalysisEDF.