Library prosa.results.edf.rta.fully_preemptive
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq path fintype bigop.
Require Import prosa.results.edf.rta.bounded_nps.
Require Export prosa.analysis.facts.preemption.task.preemptive.
Require Export prosa.analysis.facts.preemption.rtc_threshold.preemptive.
Require Export prosa.analysis.facts.readiness.basic.
Require Import prosa.results.edf.rta.bounded_nps.
Require Export prosa.analysis.facts.preemption.task.preemptive.
Require Export prosa.analysis.facts.preemption.rtc_threshold.preemptive.
Require Export prosa.analysis.facts.readiness.basic.
RTA for Fully Preemptive EDF
In this section we prove the RTA theorem for the fully preemptive EDF model
Require Import prosa.model.priority.edf.
Require Import prosa.model.processor.ideal.
Require Import prosa.model.readiness.basic.
Require Import prosa.model.processor.ideal.
Require Import prosa.model.readiness.basic.
Furthermore, we assume the fully preemptive task model. 
Consider any type of tasks ... 
 ... and any type of jobs associated with these tasks. 
  Context {Job : JobType}.
Context `{JobTask Job Task}.
Context `{JobArrival Job}.
Context `{JobCost Job}.
Context `{JobTask Job Task}.
Context `{JobArrival Job}.
Context `{JobCost Job}.
Consider any arrival sequence with consistent, non-duplicate arrivals. 
  Variable arr_seq : arrival_sequence Job.
Hypothesis H_arrival_times_are_consistent : consistent_arrival_times arr_seq.
Hypothesis H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq.
Hypothesis H_arrival_times_are_consistent : consistent_arrival_times arr_seq.
Hypothesis H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq.
Consider an arbitrary task set ts, ... 
... assume that all jobs come from this task set, ... 
... and the cost of a job cannot be larger than the task cost. 
Let max_arrivals be a family of valid arrival curves, i.e., for
     any task tsk in ts max_arrival tsk is (1) an arrival bound of
     tsk, and (2) it is a monotonic function that equals 0 for the
     empty interval delta = 0. 
  Context `{MaxArrivals Task}.
Hypothesis H_valid_arrival_curve : valid_taskset_arrival_curve ts max_arrivals.
Hypothesis H_is_arrival_curve : taskset_respects_max_arrivals arr_seq ts.
Hypothesis H_valid_arrival_curve : valid_taskset_arrival_curve ts max_arrivals.
Hypothesis H_is_arrival_curve : taskset_respects_max_arrivals arr_seq ts.
Let tsk be any task in ts that is to be analyzed. 
Next, consider any valid ideal uniprocessor schedule of the arrival sequence ... 
  Variable sched : schedule (ideal.processor_state Job).
Hypothesis H_sched_valid: valid_schedule sched arr_seq.
Hypothesis H_jobs_come_from_arrival_sequence:
jobs_come_from_arrival_sequence sched arr_seq.
Hypothesis H_sched_valid: valid_schedule sched arr_seq.
Hypothesis H_jobs_come_from_arrival_sequence:
jobs_come_from_arrival_sequence sched arr_seq.
Next, we assume that the schedule is a work-conserving schedule... 
... and the schedule respects the policy defined by the
      job_preemptable function (i.e., jobs have bounded non-preemptive
      segments). 
Total Workload and Length of Busy Interval
Using the sum of individual request bound functions, we define the request bound 
      function of all tasks (total request bound function). 
Next, we define an upper bound on interfering workload received from jobs 
      of other tasks with higher-than-or-equal priority. 
  Let bound_on_total_hep_workload A Δ :=
\sum_(tsk_o <- ts | tsk_o != tsk)
rbf tsk_o (minn ((A + ε) + task_deadline tsk - task_deadline tsk_o) Δ).
\sum_(tsk_o <- ts | tsk_o != tsk)
rbf tsk_o (minn ((A + ε) + task_deadline tsk - task_deadline tsk_o) Δ).
Let L be any positive fixed point of the busy interval recurrence. 
Response-Time Bound
Consider any value R, and assume that for any given arrival
      offset A in the search space, there is a solution of the
      response-time bound recurrence which is bounded by R. 
  Variable R : duration.
Hypothesis H_R_is_maximum:
∀ (A : duration),
is_in_search_space A →
∃ (F : duration),
A + F ≥ task_rbf (A + ε) + bound_on_total_hep_workload A (A + F) ∧
R ≥ F.
Hypothesis H_R_is_maximum:
∀ (A : duration),
is_in_search_space A →
∃ (F : duration),
A + F ≥ task_rbf (A + ε) + bound_on_total_hep_workload A (A + F) ∧
R ≥ F.
Now, we can leverage the results for the abstract model with
      bounded non-preemptive segments to establish a response-time
      bound for the more concrete model of fully preemptive
      scheduling.