Library rt.analysis.parallel.bertogna_fp_theory

Require Import rt.util.all.
Require Import rt.model.basic.task rt.model.basic.job rt.model.basic.task_arrival
               rt.model.basic.schedule rt.model.basic.platform rt.model.basic.platform_fp
               rt.model.basic.workload rt.model.basic.schedulability rt.model.basic.priority
               rt.model.basic.response_time rt.model.basic.interference.
Require Import rt.analysis.parallel.workload_bound rt.analysis.parallel.interference_bound_fp.

Module ResponseTimeAnalysisFP.

  Export Job SporadicTaskset ScheduleOfSporadicTask Workload Interference InterferenceBoundFP
         Platform PlatformFP Schedulability ResponseTime Priority SporadicTaskArrival WorkloadBound.

  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 any schedule such that...*)
    Variable num_cpus: nat.
    Variable sched: schedule num_cpus arr_seq.

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

    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 there exists at least one processor. *)
    Hypothesis H_at_least_one_cpu :
      num_cpus > 0.

    (* Assume that we have a task set (with no duplicates) where all jobs
       come from the task set and all tasks have valid parameters and constrained deadlines. *)

    Variable ts: taskset_of sporadic_task.
    Hypothesis H_ts_is_a_set: uniq ts.
    Hypothesis H_all_jobs_from_taskset:
       (j: JobIn arr_seq), job_task j \in ts.
    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.

    (* Next, consider a task tsk that is to be analyzed. *)
    Variable tsk: sporadic_task.
    Hypothesis task_in_ts: tsk \in ts.

    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.

    (* Assume a known response-time bound for any interfering task *)
    Let task_with_response_time := (sporadic_task × time)%type.
    Variable hp_bounds: seq task_with_response_time.

    (* For FP scheduling, assume there exists a fixed task priority. *)
    Variable higher_eq_priority: FP_policy sporadic_task.

    Let can_interfere_with_tsk := fp_can_interfere_with higher_eq_priority tsk.

    (* Assume that hp_bounds has exactly the tasks that interfere with tsk,... *)
    Hypothesis H_hp_bounds_has_interfering_tasks:
      [seq tsk_hp <- ts | can_interfere_with_tsk tsk_hp] = unzip1 hp_bounds.

    (* ...and that all values in the pairs contain valid response-time bounds *)
    Hypothesis H_response_time_of_interfering_tasks_is_known:
       hp_tsk R,
        (hp_tsk, R) \in hp_bounds
        response_time_bounded_by hp_tsk R.

    (* Assume that the scheduler is work-conserving and enforces the FP policy. *)
    Hypothesis H_work_conserving: work_conserving job_cost sched.
    Hypothesis H_enforces_FP_policy:
      enforces_FP_policy job_cost job_task sched higher_eq_priority.

    (* Let R be the fixed point of Bertogna and Cirinei's recurrence, ...*)
    Variable R: time.
    Hypothesis H_response_time_recurrence_holds :
      R = task_cost tsk +
          div_floor
            (total_interference_bound_fp task_cost task_period tsk hp_bounds R higher_eq_priority)
            num_cpus.

    (* ... 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 present some lemmas. *)
    Section Lemmas.

      (* Consider any job j of tsk. *)
      Variable j: JobIn arr_seq.
      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: JobIn arr_seq),
          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_cost job_task sched 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.

      (* Also, let ts_interf be the subset of tasks that interfere with tsk. *)
      Let ts_interf := [seq tsk_other <- ts | can_interfere_with_tsk tsk_other].

      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) \in hp_bounds.

        (* Note that tsk_other is in task set ts ...*)
        Lemma bertogna_fp_tsk_other_in_ts: tsk_other \in ts.

        (*... and interferes with task tsk. *)
        Lemma bertogna_fp_tsk_other_interferes: can_interfere_with_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 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 >= R - e_k + 1. *)

        Lemma bertogna_fp_too_much_interference : X R - task_cost tsk + 1.

        Let scheduled_task_other_than_tsk (t: time) (tsk_other: sporadic_task) :=
          task_is_scheduled job_task sched tsk_other t &&
          can_interfere_with_tsk tsk_other.

        (* 1) Next, we prove that the sum of the interference of each task is equal
              to the total interference multiplied by the number of processors. This
              holds because interference only occurs when all processors are busy. *)

        Lemma bertogna_fp_all_cpus_busy :
          \sum_(tsk_k <- ts_interf) x tsk_k = X × num_cpus.

        (* 2) Now, we prove that the Bertogna's interference bound
              is not enough to cover the sum of the "minimum" term over
              all tasks (artifact of the proof by contradiction). *)

        Lemma bertogna_fp_sum_exceeds_total_interference:
          \sum_((tsk_k, R_k) <- hp_bounds)
           x tsk_k > total_interference_bound_fp task_cost task_period tsk
                                            hp_bounds R higher_eq_priority.

        (* 3) After concluding that the sum of the minimum exceeds (R - e_i + 1),
              we prove that there exists a tuple (tsk_k, R_k) such that x_k > W_k. *)

        Lemma bertogna_fp_exists_task_that_exceeds_bound :
           tsk_k R_k,
            (tsk_k, R_k) \in hp_bounds
            x tsk_k > workload_bound tsk_k R_k.

      End DerivingContradiction.

    End Lemmas.

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

  End ResponseTimeBound.

End ResponseTimeAnalysisFP.