Library prosa.results.edf.rta.bounded_pi
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq path fintype bigop.
Require Export prosa.analysis.facts.priority.edf.
Require Export prosa.analysis.definitions.schedulability.
Require Import prosa.model.readiness.basic.
Require Import prosa.model.priority.edf.
Require Import prosa.model.task.absolute_deadline.
Require Import prosa.analysis.abstract.ideal_jlfp_rta.
Require Import prosa.analysis.facts.busy_interval.carry_in.
Require Import prosa.analysis.facts.readiness.basic.
Require Import prosa.analysis.facts.busy_interval.ideal.inequalities.
Require Export prosa.analysis.facts.priority.edf.
Require Export prosa.analysis.definitions.schedulability.
Require Import prosa.model.readiness.basic.
Require Import prosa.model.priority.edf.
Require Import prosa.model.task.absolute_deadline.
Require Import prosa.analysis.abstract.ideal_jlfp_rta.
Require Import prosa.analysis.facts.busy_interval.carry_in.
Require Import prosa.analysis.facts.readiness.basic.
Require Import prosa.analysis.facts.busy_interval.ideal.inequalities.
Abstract RTA for EDF-schedulers with Bounded Priority Inversion
In this module we instantiate the Abstract Response-Time analysis (aRTA) to EDF-schedulers for ideal uni-processor model of real-time tasks with arbitrary arrival models.
Consider any type of tasks ...
Context {Task : TaskType}.
Context `{TaskCost Task}.
Context `{TaskDeadline Task}.
Context `{TaskRunToCompletionThreshold Task}.
Context `{TaskMaxNonpreemptiveSegment Task}.
Context `{TaskCost Task}.
Context `{TaskDeadline Task}.
Context `{TaskRunToCompletionThreshold Task}.
Context `{TaskMaxNonpreemptiveSegment Task}.
... and any type of jobs associated with these tasks.
Context {Job : JobType}.
Context `{JobTask Job Task}.
Context {Arrival : JobArrival Job}.
Context {Cost : JobCost Job}.
Context `{JobPreemptable Job}.
Context `{JobTask Job Task}.
Context {Arrival : JobArrival Job}.
Context {Cost : JobCost Job}.
Context `{JobPreemptable Job}.
We assume the classic (i.e., Liu & Layland) model of readiness
without jitter or self-suspensions, wherein pending jobs are
always ready.
#[local] Existing Instance basic_ready_instance.
For clarity, let's denote the relative deadline of a task as D.
Consider the EDF policy that indicates a higher-or-equal priority relation.
Note that we do not relate the EDF policy with the scheduler. However, we
define functions for Interference and Interfering Workload that actively use
the concept of priorities.
Consider any arrival sequence with consistent, non-duplicate arrivals.
Variable arr_seq : arrival_sequence Job.
Hypothesis H_valid_arrival_sequence : valid_arrival_sequence arr_seq.
Hypothesis H_valid_arrival_sequence : valid_arrival_sequence arr_seq.
Next, consider any valid ideal uni-processor schedule of this arrival sequence
that follows the scheduling policy.
Variable sched : schedule (ideal.processor_state Job).
Hypothesis H_sched_valid : valid_schedule sched arr_seq.
Hypothesis H_respects_policy : respects_JLFP_policy_at_preemption_point arr_seq sched EDF.
Hypothesis H_sched_valid : valid_schedule sched arr_seq.
Hypothesis H_respects_policy : respects_JLFP_policy_at_preemption_point arr_seq sched EDF.
Note that we differentiate between abstract and
classical notions of work conserving schedule.
Let work_conserving_ab := definitions.work_conserving arr_seq sched.
Let work_conserving_cl := work_conserving.work_conserving arr_seq sched.
Let work_conserving_cl := work_conserving.work_conserving arr_seq sched.
We assume that the schedule is a work-conserving schedule
in the classical sense, and later prove that the hypothesis
about abstract work-conservation also holds.
Assume that a job cost cannot be larger than a task cost.
Consider an arbitrary task set ts.
Next, we assume that all jobs come from the task set.
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.
Consider a valid preemption model...
...and a valid task run-to-completion threshold function. That
is, task_rtct tsk is (1) no bigger than tsk's cost, (2) for
any job of task tsk job_rtct is bounded by task_rtct.
We introduce rbf as an abbreviation of the task request bound function,
which is defined as task_cost(T) × max_arrivals(T,Δ) for some task T.
Using the sum of individual request bound functions, we define the request bound
function of all tasks (total request bound function).
Assume that there exists a bound on the length of any priority inversion experienced
by any job of task tsk. Since we analyze only task tsk, we ignore the lengths of priority
inversions incurred by any other tasks.
Variable priority_inversion_bound: duration → duration.
Hypothesis H_priority_inversion_is_bounded:
priority_inversion_is_bounded_by
arr_seq sched tsk priority_inversion_bound.
Hypothesis H_priority_inversion_is_bounded:
priority_inversion_is_bounded_by
arr_seq sched tsk priority_inversion_bound.
Let L be any positive fixed point of the busy interval recurrence.
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 Δ : duration) :=
\sum_(tsk_o <- ts | tsk_o != tsk)
rbf tsk_o (minn ((A + ε) + D tsk - D tsk_o) Δ).
\sum_(tsk_o <- ts | tsk_o != tsk)
rbf tsk_o (minn ((A + ε) + D tsk - D tsk_o) Δ).
To reduce the time complexity of the analysis, we introduce the notion of search space for EDF.
Intuitively, this corresponds to all "interesting" arrival offsets that the job under
analysis might have with regard to the beginning of its busy-window.
In the case of the search space for EDF, we consider three conditions.
First, we ask whether task_rbf A ≠ task_rbf (A + ε).
Second, we ask whether there exists a task tsko from ts such that tsko
≠ tsk and rbf(tsko, A + D tsk - D tsko) ≠ rbf(tsko, A + ε + D tsk - D
tsko). Note that we use a slightly uncommon notation has (λ tsko ⇒ P
tskₒ) ts, which can be interpreted as follows: the task set ts contains
a task tsko such that a predicate P holds for tsko.
Definition bound_on_total_hep_workload_changes_at A :=
has (fun tsko ⇒
(tsk != tsko)
&& (rbf tsko (A + D tsk - D tsko)
!= rbf tsko ((A + ε) + D tsk - D tsko))) ts.
has (fun tsko ⇒
(tsk != tsko)
&& (rbf tsko (A + D tsk - D tsko)
!= rbf tsko ((A + ε) + D tsk - D tsko))) ts.
Definition priority_inversion_changes_at (A : duration) :=
priority_inversion_bound (A - ε) != priority_inversion_bound A.
priority_inversion_bound (A - ε) != priority_inversion_bound A.
The final search space for EDF is a set of offsets that are less than L
and where priority_inversion_bound, task_rbf, or
bound_on_total_hep_workload changes in value.
Definition is_in_search_space (A : duration) :=
(A < L) && (priority_inversion_changes_at A
|| task_rbf_changes_at A
|| bound_on_total_hep_workload_changes_at A).
(A < L) && (priority_inversion_changes_at A
|| task_rbf_changes_at A
|| bound_on_total_hep_workload_changes_at A).
Let R be a value that upper-bounds the solution of each
response-time recurrence, i.e., for any relative arrival time A
in the search space, there exists a corresponding solution F
such that R ≥ F + (task cost - task lock-in service).
Variable R : duration.
Hypothesis H_R_is_maximum:
∀ (A : duration),
is_in_search_space A →
∃ (F : duration),
A + F ≥ priority_inversion_bound A
+ (task_rbf (A + ε) - (task_cost tsk - task_rtct tsk))
+ bound_on_total_hep_workload A (A + F) ∧
R ≥ F + (task_cost tsk - task_rtct tsk).
Hypothesis H_R_is_maximum:
∀ (A : duration),
is_in_search_space A →
∃ (F : duration),
A + F ≥ priority_inversion_bound A
+ (task_rbf (A + ε) - (task_cost tsk - task_rtct tsk))
+ bound_on_total_hep_workload A (A + F) ∧
R ≥ F + (task_cost tsk - task_rtct tsk).
To use the theorem uniprocessor_response_time_bound_seq from the Abstract RTA module,
we need to specify functions of interference, interfering workload and IBF_other.
Instantiation of Interference We say that job j incurs interference at time t iff it cannot execute due to
a higher-or-equal-priority job being scheduled, or if it incurs a priority inversion.
Instantiation of Interfering Workload The interfering workload, in turn, is defined as the sum of the priority inversion
function and interfering workload of jobs with higher or equal priority.
Let interfering_workload (j : Job) (t : instant) :=
ideal_jlfp_rta.interfering_workload arr_seq sched j t.
ideal_jlfp_rta.interfering_workload arr_seq sched j t.
Finally, we define the interference bound function (IBF_other). IBF_other bounds
the interference if tasks are sequential. Since tasks are sequential, we exclude
interference from other jobs of the same task. For EDF, we define IBF_other as
the sum of the priority interference bound and the higher-or-equal-priority workload.
Filling Out Hypothesis Of Abstract RTA Theorem
In this section we prove that all hypotheses necessary to use the abstract theorem are satisfied.
First, we prove that IBF_other is indeed an interference bound.
Variable j : Job.
Hypothesis H_j_arrives : arrives_in arr_seq j.
Hypothesis H_job_of_tsk : job_of_task tsk j.
Hypothesis H_job_cost_positive: job_cost_positive j.
Hypothesis H_j_arrives : arrives_in arr_seq j.
Hypothesis H_job_of_tsk : job_of_task tsk j.
Hypothesis H_job_cost_positive: job_cost_positive j.
Consider any busy interval
[t1, t2)
of job j.
Variable t1 t2 : duration.
Hypothesis H_busy_interval :
definitions.busy_interval sched interference interfering_workload j t1 t2.
Hypothesis H_busy_interval :
definitions.busy_interval sched interference interfering_workload j t1 t2.
Let's define A as a relative arrival time of job j (with respect to time t1).
Consider an arbitrary shift Δ inside the busy interval ...
We define a predicate EDF_from tsk.Predicate EDF_from tsk
holds true for any job jo of task tsk such that
job_deadline jo ≤ job_deadline j.
Then we prove that the total workload of jobs with
higher-or-equal priority from task tsk_o over time
interval t1, t1 + Δ is bounded by workload over time
interval t1, t1 + A + ε + D tsk - D tsk_o.
The intuition behind this inequality is that jobs which arrive
after time instant t1 + A + ε + D tsk - D tsk_o has smaller priority
than job j due to the term D tsk - D tsk_o.
Lemma total_workload_shorten_range:
workload_of_jobs (EDF_from tsk_o) (arrivals_between arr_seq t1 (t1 + Δ))
≤ workload_of_jobs (EDF_from tsk_o) (arrivals_between arr_seq t1 (t1 + (A + ε + D tsk - D tsk_o))).
End ShortenRange.
workload_of_jobs (EDF_from tsk_o) (arrivals_between arr_seq t1 (t1 + Δ))
≤ workload_of_jobs (EDF_from tsk_o) (arrivals_between arr_seq t1 (t1 + (A + ε + D tsk - D tsk_o))).
End ShortenRange.
Using the above lemma, we prove that total
workload of tasks is at most bound_on_total_hep_workload(A, Δ).
Corollary sum_of_workloads_is_at_most_bound_on_total_hep_workload :
\sum_(tsk_o <- ts | tsk_o != tsk) workload_of_jobs (EDF_from tsk_o) jobs
≤ bound_on_total_hep_workload A Δ.
End HepWorkloadBound.
\sum_(tsk_o <- ts | tsk_o != tsk) workload_of_jobs (EDF_from tsk_o) jobs
≤ bound_on_total_hep_workload A Δ.
End HepWorkloadBound.
Recall that in module abstract_seq_RTA hypothesis
task_interference_is_bounded_by expects to receive a function
that maps some task t, the relative arrival time of a job j of
task t, and the length of the interval to the maximum amount
of interference.
However, in this module we analyze only one task -- tsk,
therefore it is “hard-coded” inside the interference bound
function IBF_other. Therefore, in order for the IBF_other signature to
match the required signature in module abstract_seq_RTA, we
wrap the IBF_other function in a function that accepts, but simply
ignores the task.
Corollary instantiated_task_interference_is_bounded:
task_interference_is_bounded_by
arr_seq sched tsk interference interfering_workload (fun tsk A R ⇒ IBF_other A R).
End TaskInterferenceIsBoundedByIBF_other.
task_interference_is_bounded_by
arr_seq sched tsk interference interfering_workload (fun tsk A R ⇒ IBF_other A R).
End TaskInterferenceIsBoundedByIBF_other.
Finally, we show that there exists a solution for the response-time recurrence.
Consider any job j of tsk.
Variable j : Job.
Hypothesis H_j_arrives : arrives_in arr_seq j.
Hypothesis H_job_of_tsk : job_of_task tsk j.
Hypothesis H_job_cost_positive : job_cost_positive j.
Hypothesis H_j_arrives : arrives_in arr_seq j.
Hypothesis H_job_of_tsk : job_of_task tsk j.
Hypothesis H_job_cost_positive : job_cost_positive j.
Given any job j of task tsk that arrives exactly A units after the beginning of
the busy interval, the bound of the total interference incurred by j within an
interval of length Δ is equal to task_rbf (A + ε) - task_cost tsk + IBF_other(A, Δ).
Let total_interference_bound tsk (A Δ : duration) :=
task_rbf (A + ε) - task_cost tsk + IBF_other A Δ.
task_rbf (A + ε) - task_cost tsk + IBF_other A Δ.
Next, consider any A from the search space (in abstract sense).
Variable A : duration.
Hypothesis H_A_is_in_abstract_search_space:
search_space.is_in_search_space tsk L total_interference_bound A.
Hypothesis H_A_is_in_abstract_search_space:
search_space.is_in_search_space tsk L total_interference_bound A.
We prove that A is also in the concrete search space.
Then, there exists solution for response-time recurrence (in the abstract sense).
Corollary correct_search_space:
∃ F,
A + F ≥ task_rbf (A + ε) - (task_cost tsk - task_rtct tsk) + IBF_other A (A + F) ∧
R ≥ F + (task_cost tsk - task_rtct tsk).
End SolutionOfResponseTimeReccurenceExists.
End FillingOutHypothesesOfAbstractRTATheorem.
∃ F,
A + F ≥ task_rbf (A + ε) - (task_cost tsk - task_rtct tsk) + IBF_other A (A + F) ∧
R ≥ F + (task_cost tsk - task_rtct tsk).
End SolutionOfResponseTimeReccurenceExists.
End FillingOutHypothesesOfAbstractRTATheorem.