Library prosa.analysis.abstract.abstract_rta
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq path fintype bigop.
Require Export prosa.analysis.definitions.schedulability.
Require Export prosa.analysis.abstract.search_space.
Require Export prosa.analysis.abstract.run_to_completion.
Require Export prosa.analysis.definitions.schedulability.
Require Export prosa.analysis.abstract.search_space.
Require Export prosa.analysis.abstract.run_to_completion.
Abstract Response-Time Analysis
In this module, we propose the general framework for response-time analysis (RTA) of uni-processor scheduling of real-time tasks with arbitrary arrival models. We prove that the maximum (with respect to the set of offsets) among the solutions of the response-time bound recurrence is a response time bound for tsk. Note that in this section we do not rely on any hypotheses about job sequentiality.
Consider any type of tasks ...
... and any type of jobs associated with these tasks.
Context {Job : JobType}.
Context `{JobTask Job Task}.
Context {JA : JobArrival Job}.
Context {JC : JobCost Job}.
Context `{JobPreemptable Job}.
Context `{JobTask Job Task}.
Context {JA : JobArrival Job}.
Context {JC : JobCost Job}.
Context `{JobPreemptable Job}.
Consider any kind of uni-service ideal processor state model.
Context {PState : ProcessorState Job}.
Hypothesis H_ideal_progress_proc_model : ideal_progress_proc_model PState.
Hypothesis H_unit_service_proc_model : unit_service_proc_model PState.
Hypothesis H_ideal_progress_proc_model : ideal_progress_proc_model PState.
Hypothesis H_unit_service_proc_model : unit_service_proc_model PState.
Consider any valid 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.
... and any schedule of this arrival sequence...
... where jobs do not execute before their arrival nor after completion.
Hypothesis H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched.
Hypothesis H_completed_jobs_dont_execute : completed_jobs_dont_execute sched.
Hypothesis H_completed_jobs_dont_execute : completed_jobs_dont_execute sched.
Assume that the job costs are no larger than the task costs.
Consider a task set ts...
... and a task tsk of 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.
Let's define some local names for clarity.
Let work_conserving := work_conserving arr_seq sched.
Let busy_intervals_are_bounded_by := busy_intervals_are_bounded_by arr_seq sched tsk.
Let job_interference_is_bounded_by := job_interference_is_bounded_by arr_seq sched tsk.
Let busy_intervals_are_bounded_by := busy_intervals_are_bounded_by arr_seq sched tsk.
Let job_interference_is_bounded_by := job_interference_is_bounded_by arr_seq sched tsk.
Assume we are provided with abstract functions for interference and interfering workload.
Variable interference : Job → instant → bool.
Variable interfering_workload : Job → instant → duration.
Variable interfering_workload : Job → instant → duration.
We assume that the scheduler is work-conserving.
For simplicity, let's define some local names.
Let cumul_interference := cumul_interference interference.
Let cumul_interfering_workload := cumul_interfering_workload interfering_workload.
Let busy_interval := busy_interval sched interference interfering_workload.
Let response_time_bounded_by := task_response_time_bound arr_seq sched.
Let cumul_interfering_workload := cumul_interfering_workload interfering_workload.
Let busy_interval := busy_interval sched interference interfering_workload.
Let response_time_bounded_by := task_response_time_bound arr_seq sched.
Let L be a constant which bounds any busy interval of task tsk.
Variable L : duration.
Hypothesis H_busy_interval_exists:
busy_intervals_are_bounded_by interference interfering_workload L.
Hypothesis H_busy_interval_exists:
busy_intervals_are_bounded_by interference interfering_workload L.
Next, assume that interference_bound_function is a bound on
the interference incurred by jobs of task tsk.
Variable interference_bound_function : Task → duration → duration → duration.
Hypothesis H_job_interference_is_bounded:
job_interference_is_bounded_by
interference interfering_workload interference_bound_function.
Hypothesis H_job_interference_is_bounded:
job_interference_is_bounded_by
interference interfering_workload interference_bound_function.
For simplicity, let's define a local name for the search space.
Consider any value R 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 tsk - task_rtct tsk).
Variable R: nat.
Hypothesis H_R_is_maximum:
∀ A,
is_in_search_space A →
∃ F,
A + F ≥ task_rtct tsk
+ interference_bound_function tsk A (A + F) ∧
R ≥ F + (task_cost tsk - task_rtct tsk).
Hypothesis H_R_is_maximum:
∀ A,
is_in_search_space A →
∃ F,
A + F ≥ task_rtct tsk
+ interference_bound_function tsk A (A + F) ∧
R ≥ F + (task_cost tsk - task_rtct tsk).
In this section we show a detailed proof of the main theorem
that establishes that R is a response-time bound of task tsk.
Consider any job j of tsk with positive cost.
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.
Assume we have a busy interval
[t1, t2)
of job j that is bounded by L.
Let's define A as a relative arrival time of job j (with respect to time t1).
In order to prove that R is a response-time bound of job j, we use hypothesis H_R_is_maximum.
Note that the relative arrival time (A) is not necessarily from the search space. However,
earlier we have proven that for any A there exists another A_sp from the search space that
shares the same IBF value. Moreover, we've also shown that there exists an F_sp such that
F_sp is a solution of the response time recurrence for parameter A_sp. Thus, despite the
fact that the relative arrival time may not lie in the search space, we can still use
the assumption H_R_is_maximum.
More formally, consider any A_sp and F_sp such that:..
(b) interference_bound_function(A, x) is equal to
interference_bound_function(A_sp, x) for all x less than L...
Hypothesis H_equivalent :
are_equivalent_at_values_less_than
(interference_bound_function tsk A)
(interference_bound_function tsk A_sp) L.
are_equivalent_at_values_less_than
(interference_bound_function tsk A)
(interference_bound_function tsk A_sp) L.
(c) A_sp is in the search space, ...
Hypothesis H_Asp_Fsp_fixpoint :
A_sp + F_sp ≥ task_rtct tsk + interference_bound_function tsk A_sp (A_sp + F_sp).
A_sp + F_sp ≥ task_rtct tsk + interference_bound_function tsk A_sp (A_sp + F_sp).
In this section, we consider the case where the solution is so large
that the value of t1 + A_sp + F_sp goes beyond the busy interval.
Although this case may be impossible in some scenarios, it can be
easily proven, since any job that completes by the end of the busy
interval remains completed.
But since we know that the job is completed by the end of its busy interval,
we can show that job j is completed by job arrival j + R.
Lemma job_completed_by_arrival_plus_R_1:
completed_by sched j (job_arrival j + R).
End FixpointOutsideBusyInterval.
completed_by sched j (job_arrival j + R).
End FixpointOutsideBusyInterval.
In this section, we consider the complementary case where
t1 + A_sp + F_sp lies inside the busy interval.
Next, let's consider two other cases: CASE 1: the value of the fix-point is no less than the relative arrival time of job j.
In this section, we prove that the fact that job j is not completed by
time job_arrival j + R leads to a contradiction. Which in turn implies
that the opposite is true -- job j completes by time job_arrival j + R.
Recall that by lemma "solution_for_A_exists" there is a solution F
of the response-time recurrence for the given relative arrival time A
(which is not necessarily from the search space).
Thus, consider a constant F such that:..
Some additional reasoning is required since the term task_cost tsk - task_rtct tsk
does not necessarily bound the term job_cost j - job_rtct j. That is, a job can
have a small run-to-completion threshold, thereby becoming non-preemptive much earlier than guaranteed
according to task run-to-completion threshold, while simultaneously executing the last non-preemptive
segment that is longer than task_cost tsk - task_rtct tsk (e.g., this is possible
in the case of floating non-preemptive sections).
In this case we cannot directly apply lemma "j_receives_at_least_run_to_completion_threshold". Therefore
we introduce two temporal notions of the last non-preemptive region of job j and an execution
optimism. We use these notions inside this proof, so we define them only locally.
Let the last non-preemptive region of job j (last) be
the difference between the cost of the job and the j's
run-to-completion threshold (i.e. job_cost j - job_rtct j).
We know that after j has reached its
run-to-completion threshold, it will additionally be
executed job_last j units of time.
And let execution optimism (optimism) be the difference
between the tsk's run-to-completion threshold and the
j's run-to-completion threshold (i.e. task_rtct -
job_rtct). Intuitively, optimism is how much earlier
job j has received its run-to-completion threshold than
it could at worst.
From lemma "j_receives_at_least_run_to_completion_threshold"
with parameters progress_of_job := job_rtct j and delta :=
(A + F) - optimism) we know that service of j by time
t1 + (A + F) - optimism is no less than job_rtct
j. Hence, job j is completed by time t1 + (A + F) -
optimism + last.
Lemma j_is_completed_by_t1_A_F_optimist_last :
completed_by sched j (t1 + (A + F - optimism) + job_last).
completed_by sched j (t1 + (A + F - optimism) + job_last).
However, t1 + (A + F) - optimism + last ≤ job_arrival j + R!
To prove this fact we need a few auxiliary inequalities that are
needed because we use the truncated subtraction in our development.
So, for example a + (b - c) = a + b - c only if b ≥ c.
Recall that we consider a busy interval of a job j, and j has arrived A time units
after the beginning the busy interval. From basic properties of a busy interval it
follows that job j incurs interference at any time instant t ∈
[t1, t1 + A)
.
Therefore interference_bound_function(tsk, A, A + F) is at least A.
... and optimism is at most F.
Next we show that t1 + (A + F) - optimism + last is at most job_arrival j + R,
which is easy to see from the following sequence of inequalities:
t1 + (A + F) - optimism + last
≤ job_arrival j + (F - optimism) + job_last
≤ job_arrival j + (F_sp - optimism) + job_last
≤ job_arrival j + F_sp + (job_last - optimism)
≤ job_arrival j + F_sp + job_cost j - task_rtct tsk
≤ job_arrival j + F_sp + task_cost tsk - task_rtct tsk
≤ job_arrival j + R.
Lemma job_completed_by_arrival_plus_R_2:
completed_by sched j (job_arrival j + R).
End FixpointIsNoLessThanArrival.
completed_by sched j (job_arrival j + R).
End FixpointIsNoLessThanArrival.
CASE 2: the value of the fix-point is less than the relative arrival time of
job j (which turns out to be impossible, i.e. the solution of the response-time
recurrence is always equal to or greater than the relative arrival time).
Note that the relative arrival time of job j is less than L.
We can use j_receives_at_least_run_to_completion_threshold to prove that the service
received by j by time t1 + (A_sp + F_sp) is no less than run-to-completion threshold.
Lemma service_of_job_ge_run_to_completion_threshold:
service sched j (t1 + (A_sp + F_sp)) ≥ job_rtct j.
service sched j (t1 + (A_sp + F_sp)) ≥ job_rtct j.
However, this is a contradiction. Since job j has not yet arrived, its service
is equal to 0. However, run-to-completion threshold is always positive.
Lemma relative_arrival_time_is_no_less_than_fixpoint:
False.
End FixpointCannotBeSmallerThanArrival.
End FixpointInsideBusyInterval.
End ProofOfTheorem.
False.
End FixpointCannotBeSmallerThanArrival.
End FixpointInsideBusyInterval.
End ProofOfTheorem.
Using the lemmas above, we prove that R is a response-time bound.