Library prosa.classic.analysis.uni.basic.workload_bound_fp

Require Import prosa.classic.util.all.
Require Import prosa.classic.model.arrival.basic.task prosa.classic.model.arrival.basic.job prosa.classic.model.priority
               prosa.classic.model.arrival.basic.task_arrival prosa.classic.model.arrival.basic.arrival_bounds.
Require Import prosa.classic.model.schedule.uni.schedule prosa.classic.model.schedule.uni.workload.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq fintype bigop div.

Module WorkloadBoundFP.

  Import Job SporadicTaskset UniprocessorSchedule Priority Workload
         TaskArrival ArrivalBounds.

  (* In this section, we define a bound for the workload of a single task
     under uniprocessor FP scheduling. *)

  Section SingleTask.

    Context {Task: eqType}.
    Variable task_cost: Task time.
    Variable task_period: Task time.

    (* Consider any task tsk that is to be scheduled in an interval of length delta. *)
    Variable tsk: Task.
    Variable delta: time.

    (* Based on the maximum number of jobs of tsk that can execute in the interval, ... *)
    Definition max_jobs := div_ceil delta (task_period tsk).

    (* ... we define the following workload bound for the task. *)
    Definition task_workload_bound_FP := max_jobs × task_cost tsk.

  End SingleTask.

  (* In this section, we define a bound for the workload of multiple tasks. *)
  Section AllTasks.

    Context {Task: eqType}.
    Variable task_cost: Task time.
    Variable task_period: Task time.

    (* Assume any FP policy. *)
    Variable higher_eq_priority: FP_policy Task.

    (* Consider a task set ts... *)
    Variable ts: list Task.

    (* ...and let tsk be the task to be analyzed. *)
    Variable tsk: Task.

    (* Let delta be the length of the interval of interest. *)
    Variable delta: time.

    (* Recall the definition of higher-or-equal-priority task and
       the per-task workload bound for FP scheduling. *)

    Let is_hep_task tsk_other := higher_eq_priority tsk_other tsk.
    Let W tsk_other :=
      task_workload_bound_FP task_cost task_period tsk_other delta.

    (* Using the sum of individual workload bounds, we define the following bound
       for the total workload of tasks of higher-or-equal priority (with respect
       to tsk) in any interval of length delta. *)

    Definition total_workload_bound_fp :=
      \sum_(tsk_other <- ts | is_hep_task tsk_other) W tsk_other.

  End AllTasks.

  (* In this section, we prove some basic lemmas about the workload bound. *)
  Section BasicLemmas.

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

    (* Assume any FP policy. *)
    Variable higher_eq_priority: FP_policy Task.

    (* Consider a task set ts... *)
    Variable ts: list Task.

    (* ...and let tsk be any task in ts. *)
    Variable tsk: Task.
    Hypothesis H_tsk_in_ts: tsk \in ts.

    (* Recall the workload bound for uniprocessor FP scheduling. *)
    Let workload_bound :=
      total_workload_bound_fp task_cost task_period higher_eq_priority ts tsk.

    (* In this section we prove that the workload bound in a time window of
       length (task_cost tsk) is as large as (task_cost tsk) time units.
       (This is an important initial condition for the response-time analysis.) *)

    Section NoSmallerThanCost.

      (* Assume that the priority order is reflexive. *)
      Hypothesis H_priority_is_reflexive: FP_is_reflexive higher_eq_priority.

      (* Assume that cost and period of the task are positive. *)
      Hypothesis H_cost_positive: task_cost tsk > 0.
      Hypothesis H_period_positive: task_period tsk > 0.

      (* We prove that the workload bound of an interval of size (task_cost tsk)
         cannot be smaller than (task_cost tsk). *)

      Lemma total_workload_bound_fp_ge_cost:
        workload_bound (task_cost tsk) task_cost tsk.

    End NoSmallerThanCost.

    (* In this section, we prove that the workload bound is monotonically non-decreasing. *)
    Section NonDecreasing.

      (* Assume that the period of every task in the task set is positive. *)
      Hypothesis H_period_positive:
         tsk,
          tsk \in ts
          task_period tsk > 0.

      (* Then, the workload bound is a monotonically non-decreasing function.
         (This property is important for the fixed-point iteration.) *)

      Lemma total_workload_bound_fp_non_decreasing:
         delta1 delta2,
          delta1 delta2
          workload_bound delta1 workload_bound delta2.

    End NonDecreasing.

  End BasicLemmas.

  (* In this section, we prove that any fixed point R = workload_bound R
     is indeed a workload bound for an interval of length R. *)

  Section ProofWorkloadBound.

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

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

    (* Let ts be any task set with valid task parameters. *)
    Variable ts: seq Task.
    Hypothesis H_valid_task_parameters:
      valid_sporadic_taskset task_cost task_period task_deadline ts.

    (* Consider any job arrival sequence with consistent, duplicate-free arrivals. *)
    Variable arr_seq: arrival_sequence Job.
    Hypothesis H_arrival_times_are_consistent: arrival_times_are_consistent job_arrival arr_seq.
    Hypothesis H_arr_seq_is_a_set: arrival_sequence_is_a_set arr_seq.

    (* Assume that all jobs come from the task set ...*)
    Hypothesis H_all_jobs_from_taskset:
       j, arrives_in arr_seq j job_task j \in ts.

    (* ...and have valid parameters. *)
    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 jobs arrived sporadically. *)
    Hypothesis H_sporadic_arrivals:
      sporadic_task_model task_period job_arrival job_task arr_seq.

    (* Let tsk be any task in ts. *)
    Variable tsk: Task.
    Hypothesis H_tsk_in_ts: tsk \in ts.

    (* Assume any fixed-priority policy. *)
    Variable higher_eq_priority: FP_policy Task.

    (* First, let's define some local names for clarity. *)
    Let arrivals_between := jobs_arrived_between arr_seq.
    Let hp_workload t1 t2:=
      workload_of_higher_or_equal_priority_tasks job_cost job_task (arrivals_between t1 t2)
                                                 higher_eq_priority tsk.
    Let workload_bound :=
      total_workload_bound_fp task_cost task_period higher_eq_priority ts tsk.

    (* Consider any R that is a fixed point of the following equation,
       i.e., the claimed workload bound is equal to the interval length. *)

    Variable R: time.
    Hypothesis H_fixed_point: R = workload_bound R.

    (* Then, we prove that R is indeed a workload bound. *)
    Lemma fp_workload_bound_holds:
       t,
        hp_workload t (t + R) R.

  End ProofWorkloadBound.

End WorkloadBoundFP.