Library prosa.results.gel.rta.bounded_pi
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq path fintype bigop ssrZ.
Require Import prosa.model.readiness.basic.
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.model.schedule.priority_driven.
Require Import prosa.model.priority.gel.
Require Import analysis.facts.priority.gel.
Require Import prosa.analysis.facts.busy_interval.ideal.inequalities.
Require Import prosa.analysis.facts.model.workload.
Require Import prosa.model.readiness.basic.
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.model.schedule.priority_driven.
Require Import prosa.model.priority.gel.
Require Import analysis.facts.priority.gel.
Require Import prosa.analysis.facts.busy_interval.ideal.inequalities.
Require Import prosa.analysis.facts.model.workload.
Abstract RTA for GEL-Schedulers with Bounded Priority Inversion
In this module we instantiate the abstract response-time analysis (aRTA) for GEL-schedulers assuming the ideal uni-processor model and real-time tasks with arbitrary arrival models.
We consider tasks and jobs with the given parameters.
Context {Task : TaskType} `{TaskCost Task} `{TaskRunToCompletionThreshold Task}
`{TaskMaxNonpreemptiveSegment Task} `{MaxArrivals Task}.
Context {Job : JobType} `{JobTask Job Task} {Arrival : JobArrival Job}
{Cost : JobCost Job} `{JobPreemptable Job} `{PriorityPoint Task}.
`{TaskMaxNonpreemptiveSegment Task} `{MaxArrivals Task}.
Context {Job : JobType} `{JobTask Job Task} {Arrival : JobArrival Job}
{Cost : JobCost Job} `{JobPreemptable Job} `{PriorityPoint Task}.
We assume the basic readiness model.
#[local] Existing Instance basic_ready_instance.
A. Defining the System Model
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.
We consider a valid, ideal uniprocessor schedule...
Variable sched : schedule (ideal.processor_state Job).
Hypothesis H_sched_valid : valid_schedule sched arr_seq.
Hypothesis H_sched_valid : valid_schedule sched arr_seq.
... that respects the GEL policy at every preemption point.
Hypothesis H_respects_policy : respects_JLFP_policy_at_preemption_point arr_seq sched (GEL Job Task).
We assume that all arrivals have valid job costs.
We model the tasks in the system using a task set ts.
We assume that all jobs in the system come from this task set.
We assume that the task set respects the arrival curve
defined by max_arrivals.
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_is_arrival_curve : taskset_respects_max_arrivals arr_seq ts.
While we do not assume an explicit preemption model,
we assume that any preemption model under consideration
is valid. We also assume that the run-to-completion-threshold
of the task tsk is valid.
Hypothesis H_valid_preemption_model:
valid_preemption_model arr_seq sched.
Hypothesis H_valid_run_to_completion_threshold:
valid_task_run_to_completion_threshold arr_seq tsk.
valid_preemption_model arr_seq sched.
Hypothesis H_valid_run_to_completion_threshold:
valid_task_run_to_completion_threshold arr_seq tsk.
As mentioned, we assume that the priority inversion for the task
tsk is bound by priority_inversion_bound.
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.
B. Encoding the Scheduling Policy and Preemption Model
Next, we encode the scheduling policy and preemption model using the functions interference and interfering_workload. To this end, we simply reuse the general definitions of interference and interfering workload that apply to any JLFP policy, as defined in the module ideal_jlfp_rta.
Let interference (j : Job) (t : instant) :=
ideal_jlfp_rta.interference arr_seq sched j t.
Let interfering_workload (j : Job) (t : instant) :=
ideal_jlfp_rta.interfering_workload arr_seq sched j t.
ideal_jlfp_rta.interference arr_seq sched j t.
Let interfering_workload (j : Job) (t : instant) :=
ideal_jlfp_rta.interfering_workload arr_seq sched j t.
C. Abstract Work Conservation
Let us recall the abstract and classic notations of work conservation.
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 work-conserving in the classic sense,
which allows us to apply instantiated_i_and_w_are_coherent_with_schedule
to conclude that abstract work-conservation also holds.
D. Bounding the Maximum Busy-Window Length
Next, we define L as the fixed point of the given equation. Given this definition of L, we can apply the theorem instantiated_busy_intervals_are_bounded to prove that L bounds the length of the busy window.
Variable L : duration.
Hypothesis H_L_positive : L > 0.
Hypothesis H_fixed_point : L = total_request_bound_function ts L.
Hypothesis H_L_positive : L > 0.
Hypothesis H_fixed_point : L = total_request_bound_function ts L.
E. Defining IBF_other
Next, we define IBF_other and prove that IBF_other bounds the interference incurred by any job of tsk.
Let interval (tsk_o : Task) (A : instant):=
(Z.of_nat(A + ε) + task_priority_point tsk - task_priority_point tsk_o)%Z.
(Z.of_nat(A + ε) + task_priority_point tsk - task_priority_point tsk_o)%Z.
We define the bound on the total higher-or-equal-priority (HEP) workload
produced during the interval Δ as the sum of the RBFs of all tasks in
the task set ts (excluding tsk) over the minimum of Δ and interval.
Let bound_on_total_hep_workload (A Δ : duration) :=
\sum_(tsk_o <- ts | tsk_o != tsk)
task_request_bound_function tsk_o (minn (Z.to_nat (interval tsk_o A)) Δ).
\sum_(tsk_o <- ts | tsk_o != tsk)
task_request_bound_function tsk_o (minn (Z.to_nat (interval tsk_o A)) Δ).
Finally, IBF_other for an interval R is defined as the sum of
bound_on_total_hep_workload in R and priority_inversion_bound.
For convenience, we define the following acronym.
In this section, we prove the soundness of IBF_other. We start
by establishing a bound on the HEP workload.
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 the busy interval of j is given by
[t1,t2)
.
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.
Consider any arbitrary interval Δ within the busy window.
Consider the set of jobs arriving in
[t1, t1 + Δ)
.
We define a predicate to identify higher priority jobs coming from the task tsk.
... then the workload of jobs satisfying the predicate GEL_from
in the interval
[t1,t1 + Δ)
is equal to the workload in
the interval [t1, t1 + interval [tsk_o] [A])
.
Note that, we use the functions Z.to_nat to Z.of_nat to convert
integers to natural numbers and vice-versa.
Lemma total_workload_shorten_range:
workload_of_jobs (GEL_from tsk_o)
(arrivals_between arr_seq t1 (t1 + Δ))
≤ workload_of_jobs (GEL_from tsk_o)
(arrivals_between arr_seq t1
(Z.to_nat (Z.of_nat t1 + interval tsk_o A))).
End ShortenRange.
workload_of_jobs (GEL_from tsk_o)
(arrivals_between arr_seq t1 (t1 + Δ))
≤ workload_of_jobs (GEL_from tsk_o)
(arrivals_between arr_seq t1
(Z.to_nat (Z.of_nat t1 + interval tsk_o A))).
End ShortenRange.
Using the above lemma, we prove that bound_on_total_hep_workload
bounds the sum of higher-priority workload over all tasks in ts.
Corollary sum_of_workloads_is_at_most_bound_on_total_hep_workload :
\sum_(tsk_o <- ts | tsk_o != tsk) workload_of_jobs (GEL_from tsk_o) jobs
≤ bound_on_total_hep_workload A Δ.
End HepWorkloadBound.
\sum_(tsk_o <- ts | tsk_o != tsk) workload_of_jobs (GEL_from tsk_o) jobs
≤ bound_on_total_hep_workload A Δ.
End HepWorkloadBound.
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).
task_interference_is_bounded_by
arr_seq sched tsk interference interfering_workload (fun tsk A R ⇒ IBF_other A R).
F. Defining the Search Space
Definition task_rbf_changes_at (A : duration) :=
task_request_bound_function tsk A != task_request_bound_function tsk (A + ε).
task_request_bound_function tsk A != task_request_bound_function tsk (A + ε).
Definition bound_on_total_hep_workload_changes_at A :=
has (fun tsko ⇒
((tsk != tsko) && (interval tsko (A - ε) != interval tsko A)))
ts.
has (fun tsko ⇒
((tsk != tsko) && (interval tsko (A - ε) != interval tsko A)))
ts.
... and priority_inversion_changes_at.
Definition priority_inversion_changes_at (A : duration) :=
priority_inversion_bound (A - ε) != priority_inversion_bound A.
priority_inversion_bound (A - ε) != priority_inversion_bound A.
Finally, we define the concrete search space as the set containing
all points less than L at which any of the bounds on priority inversion, task rbf,
or bound on total HEP workload changes.
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).
In this section, we prove that for any job j of tsk,
if A is in the abstract search space, then it is also in
the concrete search space.
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.
For j, the total interference bound is defined as follows.
Let total_interference_bound tsk (A Δ : duration) :=
task_request_bound_function tsk (A + ε) - task_cost tsk + IBF_other A Δ.
task_request_bound_function tsk (A + ε) - task_cost tsk + IBF_other A Δ.
Consider any point A in the abstract search space.
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.
Then, A is also in the concrete search space.
Variable R : duration.
Hypothesis H_R_is_maximum:
∀ (A : duration),
is_in_search_space A →
∃ (F : duration),
A + F ≥ priority_inversion_bound A
+ (task_request_bound_function tsk (A + ε) - (task_cost tsk - task_rtct tsk))
+ bound_on_total_hep_workload A (A + F) ∧
R ≥ F + (task_cost tsk - task_rtct tsk).
Section ResponseTimeReccurence.
Hypothesis H_R_is_maximum:
∀ (A : duration),
is_in_search_space A →
∃ (F : duration),
A + F ≥ priority_inversion_bound A
+ (task_request_bound_function tsk (A + ε) - (task_cost tsk - task_rtct tsk))
+ bound_on_total_hep_workload A (A + F) ∧
R ≥ F + (task_cost tsk - task_rtct tsk).
Section ResponseTimeReccurence.
In order to connect the concrete definition of R
with the shape of response-time bound equation that aRTA expects,
we prove the theorem correct_search_space.
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.
We define the total interference as defined above.
Let total_interference_bound tsk (A Δ : duration) :=
task_request_bound_function tsk (A + ε) - task_cost tsk + IBF_other A Δ.
task_request_bound_function tsk (A + ε) - task_cost tsk + IBF_other A Δ.
We know that if A is in the abstract search then it is in the concrete search space.
We also know that if A is in the concrete search space then there exists an R that
satisfies H_R_is_maximum. Using these facts, here we prove that if A
is in the abstract search space then, there exists a solution to the response-time
equation as stated in the aRTA.
Corollary correct_search_space:
∀ A,
search_space.is_in_search_space tsk L total_interference_bound A →
∃ F,
A + F ≥ task_request_bound_function tsk (A + ε) - (task_cost tsk - task_rtct tsk) +
IBF_other A (A + F) ∧ R ≥ F + (task_cost tsk - task_rtct tsk).
End ResponseTimeReccurence.
∀ A,
search_space.is_in_search_space tsk L total_interference_bound A →
∃ F,
A + F ≥ task_request_bound_function tsk (A + ε) - (task_cost tsk - task_rtct tsk) +
IBF_other A (A + F) ∧ R ≥ F + (task_cost tsk - task_rtct tsk).
End ResponseTimeReccurence.