Library rt.restructuring.analysis.arrival.workload_bound

From rt.util Require Import sum.
From rt.restructuring.behavior Require Export all.
From rt.restructuring.model Require Import task schedule.priority_based.priorities.
From rt.restructuring.model.aggregate Require Import task_arrivals.
From rt.restructuring.model.arrival Require Import arrival_curves.
From rt.restructuring.analysis Require Import
     workload ideal_schedule basic_facts.arrivals.

From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq path fintype bigop.

Task Workload Bounded by Arrival Curves

Consider any type of tasks ...
  Context {Task : TaskType}.
  Context `{TaskCost Task}.

... and any type of jobs associated with these tasks.
  Context {Job : JobType}.
  Context `{JobTask Job Task}.
  Context `{JobArrival Job}.
  Context `{JobCost Job}.

Consider any arrival sequence with consistent, non-duplicate arrivals...
... and any ideal uniprocessor schedule of this arrival sequence.
Consider an FP policy that indicates a higher-or-equal priority relation.
For simplicity, let's define some local names.
We define the notion of request bound function.
  Section RequestBoundFunction.

Let MaxArrivals denote any function that takes a task and an interval length and returns the associated number of job arrivals of the task.
    Context `{MaxArrivals Task}.

In this section, we define a bound for the workload of a single task under uniprocessor FP scheduling.
    Section SingleTask.

Consider any task tsk that is to be scheduled in an interval of length delta.
      Variable tsk : Task.
      Variable delta : duration.

We define the following workload bound for the task.
      Definition task_request_bound_function := task_cost tsk × max_arrivals tsk delta.

    End SingleTask.

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

Consider a task set ts...
      Variable ts : list Task.

...and let tsk be any task in task set.
      Variable tsk : Task.

Let delta be the length of the interval of interest.
      Variable delta : duration.

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 is_other_hep_task tsk_other := higher_eq_priority tsk_other tsk && (tsk_other != tsk).

Using the sum of individual workload bounds, we define the following bound for the total workload of tasks in any interval of length delta.
      Definition total_request_bound_function :=
        \sum_(tsk <- ts) task_request_bound_function tsk delta.

Similarly, 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_hep_request_bound_function_FP :=
        \sum_(tsk_other <- ts | is_hep_task tsk_other)
         task_request_bound_function tsk_other delta.

We also define a bound for the total workload of higher-or-equal priority tasks other than tsk in any interval of length delta.
      Definition total_ohep_request_bound_function_FP :=
        \sum_(tsk_other <- ts | is_other_hep_task tsk_other)
         task_request_bound_function tsk_other delta.

    End AllTasks.

  End RequestBoundFunction.

In this section we prove some lemmas about request bound functions.
  Section ProofWorkloadBound.

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.

Assume that the job costs are no larger than the task costs.
Next, we assume that all jobs come from the task set.
Let max_arrivals be any arrival bound for taskset ts.
    Context `{MaxArrivals Task}.
    Hypothesis H_is_arrival_bound : taskset_respects_max_arrivals arr_seq ts.

Let's define some local names for clarity.
Next, we 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.

Next, we say that two jobs j1 and j2 are in relation other_higher_eq_priority, iff j1 has higher or equal priority than j2 and is produced by a different task.
Next, we recall the notions of total workload of jobs...
...notions of workload of higher or equal priority jobs...
    Let total_hep_workload t1 t2 :=
      workload_of_jobs (fun j_otherjlfp_higher_eq_priority j_other j) (arrivals_between t1 t2).

... workload of other higher or equal priority jobs...
    Let total_ohep_workload t1 t2 :=
      workload_of_jobs (fun j_otherother_higher_eq_priority j_other j) (arrivals_between t1 t2).

... and the workload of jobs of the same task as job j.
    Let task_workload t1 t2 :=
      workload_of_jobs (job_of_task tsk) (arrivals_between t1 t2).

In this section we prove that the workload of any jobs is no larger than the request bound function.
    Section WorkloadIsBoundedByRBF.

Consider any time t and any interval of length delta.
      Variable t : instant.
      Variable delta : instant.

First, we show that workload of task tsk is bounded by the number of arrivals of the task times the cost of the task.
As a corollary, we prove that workload of task is no larger the than task request bound function.
      Corollary task_workload_le_task_rbf:
        task_workload t (t + delta) task_rbf delta.

Next, we prove that total workload of other tasks with higher-or-equal priority is no larger than the total request bound function.
Next, we prove that total workload of tasks with higher-or-equal priority is no larger than the total request bound function.
Next, we prove that total workload of tasks is no larger than the total request bound function.