Library rt.analysis.apa.bertogna_fp_theory

Require Import rt.util.all.
Require Import rt.model.arrival.basic.task rt.model.arrival.basic.job rt.model.priority rt.model.arrival.basic.task_arrival.
Require Import rt.model.schedule.global.response_time rt.model.schedule.global.schedulability
               rt.model.schedule.global.workload.
Require Import rt.model.schedule.global.basic.schedule.
Require Import rt.model.schedule.apa.platform rt.model.schedule.apa.constrained_deadlines
               rt.model.schedule.apa.interference rt.model.schedule.apa.affinity.
Require Import rt.analysis.apa.workload_bound
               rt.analysis.apa.interference_bound_fp.

Module ResponseTimeAnalysisFP.

  Export Job SporadicTaskset ScheduleOfSporadicTask Workload Interference InterferenceBoundFP
         Platform Schedulability ResponseTime Priority
         TaskArrival WorkloadBound Affinity ConstrainedDeadlines.

  (* In this section, we prove that any fixed point in the APA-reduction of Bertogna
     and Cirinei's RTA for FP scheduling with slack updates is a safe response-time
     bound. This result corresponds to Lemma 9 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_arrival: Job time.
    Variable job_cost: Job time.
    Variable job_deadline: Job time.
    Variable job_task: Job sporadic_task.

    (* Assume any job arrival sequence... *)
    Variable arr_seq: arrival_sequence Job.

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

    (* Assume that we have a task set 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 \in ts task_deadline tsk task_period tsk.

    (* ... and that all jobs in the arrival sequence come from the task set. *)
    Hypothesis H_all_jobs_from_taskset:
       j,
        arrives_in arr_seq j job_task j \in 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 Job num_cpus.
    Hypothesis H_jobs_come_from_arrival_sequence:
      jobs_come_from_arrival_sequence sched 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 job_arrival sched.
    Hypothesis H_completed_jobs_dont_execute: completed_jobs_dont_execute job_cost sched.

    (* Consider a given FP policy, ... *)
    Variable higher_eq_priority: FP_policy sporadic_task.

    (* ... and assume that the schedule is an APA work-conserving
       schedule that respects this policy. *)

    Hypothesis H_respects_affinity: respects_affinity job_task sched alpha.
    Hypothesis H_work_conserving: apa_work_conserving job_arrival job_cost job_task
                                                      arr_seq sched alpha.
    Hypothesis H_respects_FP_policy:
      respects_FP_policy_under_weak_APA job_arrival job_cost job_task arr_seq
                                        sched alpha higher_eq_priority.

    (* 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_arrival job_cost job_deadline job_task arr_seq sched tsk.
    Let response_time_bounded_by (tsk: sporadic_task) :=
      is_response_time_bound_of_task job_arrival job_cost job_task arr_seq sched tsk.

    (* Next, we consider the response-time recurrence.
       Let tsk be a task in ts that is to be analyzed. *)

    Variable tsk: sporadic_task.
    Hypothesis task_in_ts: tsk \in ts.

    (* When computing the response-time bounds, we assume that each task under
       analysis has a non-empty subaffinity alpha'. *)

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

    (* Let (hp_task_in alpha') denote whether a task is a higher-priority task
       (with respect to tsk) that can execute on processors in (alpha' tsk). *)

    Let hp_task_in alpha' := higher_priority_task_in alpha higher_eq_priority tsk alpha'.

    (* Assume a response-time bound is known... *)
    Let task_with_response_time := (sporadic_task × time)%type.
    Variable hp_bounds: seq task_with_response_time.
    Hypothesis H_response_time_of_interfering_tasks_is_known:
       hp_tsk R,
        (hp_tsk, R) \in hp_bounds
        is_response_time_bound_of_task job_arrival job_cost job_task arr_seq sched hp_tsk R.

    (* ... for every higher-priority task in (alpha tsk). *)
    Hypothesis H_hp_bounds_has_interfering_tasks:
       hp_tsk,
        hp_tsk \in ts
        hp_task_in (alpha tsk) hp_tsk
         R,
          (hp_tsk, R) \in hp_bounds.

    (* Assume that the response-time bounds are larger than task costs. *)
    Hypothesis H_response_time_bounds_ge_cost:
       hp_tsk R,
        (hp_tsk, R) \in hp_bounds R task_cost hp_tsk.

    (* Assume that no deadline is missed by any higher-priority task. *)
    Hypothesis H_interfering_tasks_miss_no_deadlines:
       hp_tsk R,
        (hp_tsk, R) \in hp_bounds R task_deadline hp_tsk.

    (* Let R be the fixed point of the APA-reduction of Bertogna and
       Cirinei's response-time recurrence (using subaffinity alpha'), ...*)

    Variable R: time.
    Hypothesis H_response_time_recurrence_holds :
      R = task_cost tsk +
          div_floor
            (total_interference_bound_fp task_cost task_period alpha tsk
                            (alpha' tsk) hp_bounds R higher_eq_priority)
            #|alpha' tsk|.

    (* ... and assume that R is no larger than the deadline of tsk.*)
    Hypothesis H_response_time_no_larger_than_deadline:
      R task_deadline tsk.

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

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

      (* Assume that job j is the first job of tsk not to complete by the response time bound. *)
      Hypothesis H_j_not_completed: ~~ completed job_cost sched j (job_arrival j + R).
      Hypothesis H_previous_jobs_of_tsk_completed :
         j0,
          arrives_in arr_seq j0
          job_task j0 = tsk
          job_arrival j0 < job_arrival j
          completed job_cost sched j0 (job_arrival j0 + R).

      (* Let's call x the interference incurred by job j due to tsk_other, ...*)
      Let x (tsk_other: sporadic_task) :=
        task_interference job_arrival 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_arrival 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.

      (* Let (hp_tasks_in alpha') denote the set of higher-priority tasks
         that can be scheduled on subaffinity alpha'. *)

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

      (* Now we establish results the higher-priority tasks. *)
      Section LemmasAboutHPTasks.

        (* 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_tsk_other_already_processed: (tsk_other, R_other) \in hp_bounds.
        Hypothesis H_tsk_other_has_higher_priority: hp_task_in (alpha tsk) 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_fp_workload_bounds_interference :
          x tsk_other workload_bound tsk_other R_other.

      End LemmasAboutHPTasks.

      (* 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 >= R - e_k + 1. *)

        Lemma bertogna_fp_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_fp_interference_by_different_tasks :
           t j_other,
            job_arrival j t < job_arrival j + R
            arrives_in arr_seq j_other
            backlogged job_arrival 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 of higher-priority tasks released before the end of the
              interval complete by their periods. This follows trivially from
              the fact that all these tasks are known to be schedulable. 
              With this lemma, we can conclude that during job j's scheduling
              window there cannot be multiple pending jobs of higher-priority tasks.*)

        Lemma bertogna_fp_previous_interfering_jobs_complete_by_their_period:
           j0,
            arrives_in arr_seq j0
            hp_task_in (alpha tsk) (job_task j0)
            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_fp_all_cpus_in_affinity_busy :
          \sum_(tsk_k <- hp_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_fp_all_cpus_in_subaffinity_busy :
          \sum_(tsk_k <- hp_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_fp_alpha'_is_full:
           t,
            job_arrival j t < job_arrival j + R
            backlogged job_arrival job_cost sched j t
            count (scheduled_on_alpha_tsk t) (hp_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) (hp_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_fp_interference_in_non_full_processors :
           delta,
            0 < num_tasks_exceeding delta < #|alpha' tsk|
            \sum_(i <- hp_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_fp_minimum_exceeds_interference :
           delta,
            \sum_(tsk_k <- hp_tasks_in alpha') x tsk_k delta × #|alpha' tsk|
               \sum_(tsk_k <- hp_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_fp_interference_on_subaffinity :
           delta,
            \sum_(tsk_k <- hp_tasks_in alpha) x tsk_k delta × #|alpha tsk|
            \sum_(tsk_k <- hp_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_fp_sum_exceeds_total_interference:
          \sum_((tsk_other, R_other) <- hp_bounds | hp_task_in (alpha' tsk) tsk_other)
           minn (x tsk_other) (R - task_cost tsk + 1) >
          total_interference_bound_fp task_cost task_period alpha tsk
                           (alpha' tsk) hp_bounds R higher_eq_priority.

        (* 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, R - e_i + 1).
              This implies that x_k > W_k, which is of course a contradiction,
              since W_k is a valid task interference bound. *)

        Lemma bertogna_fp_exists_task_that_exceeds_bound :
           tsk_k R_k,
            (tsk_k, R_k) \in hp_bounds
            (minn (x tsk_k) (R - task_cost tsk + 1) >
              minn (workload_bound tsk_k R_k) (R - task_cost tsk + 1)).

      End DerivingContradiction.

    End Lemmas.

    (* Using the lemmas above, we prove that R bounds the response time of tsk. *)
    Theorem bertogna_cirinei_response_time_bound_fp :
      response_time_bounded_by tsk R.

  End ResponseTimeBound.

End ResponseTimeAnalysisFP.