Library prosa.results.edf.rta.bounded_nps
Require Import prosa.model.readiness.basic.
Require Import prosa.model.priority.edf.
Require Import prosa.model.schedule.work_conserving.
Require Import prosa.model.task.preemption.parameters.
Require Export prosa.analysis.facts.model.rbf.
Require Export prosa.analysis.facts.model.arrival_curves.
Require Export prosa.analysis.facts.model.sequential.
Require Export prosa.analysis.facts.busy_interval.pi_bound.
Require Export prosa.analysis.facts.busy_interval.arrival.
Require Export prosa.results.edf.rta.bounded_pi.
Require Export prosa.util.tactics.
Require Import prosa.model.priority.edf.
Require Import prosa.model.schedule.work_conserving.
Require Import prosa.model.task.preemption.parameters.
Require Export prosa.analysis.facts.model.rbf.
Require Export prosa.analysis.facts.model.arrival_curves.
Require Export prosa.analysis.facts.model.sequential.
Require Export prosa.analysis.facts.busy_interval.pi_bound.
Require Export prosa.analysis.facts.busy_interval.arrival.
Require Export prosa.results.edf.rta.bounded_pi.
Require Export prosa.util.tactics.
RTA for EDF with Bounded Non-Preemptive Segments
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 `{JobTask Job Task}.
Context `{Arrival : JobArrival Job}.
Context `{Cost : JobCost Job}.
We assume the classic (i.e., Liu & Layland) model of readiness
without jitter or self-suspensions, wherein pending jobs are
always ready.
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 ...
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.
In addition, we assume the existence of a function mapping jobs
to their preemption points ...
... and assume that it defines a valid preemption model with
bounded non-preemptive segments.
Hypothesis H_valid_model_with_bounded_nonpreemptive_segments:
valid_model_with_bounded_nonpreemptive_segments arr_seq sched.
valid_model_with_bounded_nonpreemptive_segments arr_seq sched.
Next, we assume that the schedule is a work-conserving schedule...
... and the schedule respects the scheduling policy at every preemption point.
Consider an arbitrary task set ts, ...
... assume that all jobs come from the 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.
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 as an abbreviation rbf for the task request bound function,
which is defined as task_cost(T) × max_arrivals(T,Δ) for a task T.
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 + ε) + D tsk - D tsk_o) Δ).
\sum_(tsk_o <- ts | tsk_o != tsk)
rbf tsk_o (minn ((A + ε) + D tsk - D tsk_o) Δ).
Let's define some local names for clarity.
Let max_length_of_priority_inversion :=
max_length_of_priority_inversion arr_seq.
Let response_time_bounded_by := task_response_time_bound arr_seq sched.
max_length_of_priority_inversion arr_seq.
Let response_time_bounded_by := task_response_time_bound arr_seq sched.
For a job with the relative arrival offset A within its busy window, we
define the following blocking bound. Only other tasks that potentially
release non-zero-cost jobs are relevant, so we define a predicate to
exclude pathological cases.
Definition blocking_relevant (tsk_o : Task) :=
(max_arrivals tsk_o ε > 0) && (task_cost tsk_o > 0).
Definition blocking_bound (A : duration) :=
\max_(tsk_o <- ts | blocking_relevant tsk_o && (D tsk_o > D tsk + A))
(task_max_nonpreemptive_segment tsk_o - ε).
(max_arrivals tsk_o ε > 0) && (task_cost tsk_o > 0).
Definition blocking_bound (A : duration) :=
\max_(tsk_o <- ts | blocking_relevant tsk_o && (D tsk_o > D tsk + A))
(task_max_nonpreemptive_segment tsk_o - ε).
Search Space
Definition is_in_search_space (L A : duration) :=
(A < L) && (task_rbf_changes_at tsk A
|| bound_on_total_hep_workload_changes_at ts tsk A).
(A < L) && (task_rbf_changes_at tsk A
|| bound_on_total_hep_workload_changes_at ts tsk A).
For the following proof, we exploit the fact that the blocking bound is
monotonically decreasing in A, which we note here.
Fact blocking_bound_decreasing :
∀ A1 A2,
A1 ≤ A2 →
blocking_bound A1 ≥ blocking_bound A2.
Proof.
move⇒ A1 A2 LEQ.
rewrite /blocking_bound.
apply: bigmax_subset ⇒ tsk_o IN /andP[/andP[OTHER LT] ARR].
by repeat (apply /andP; split) ⇒ //; lia.
Qed.
∀ A1 A2,
A1 ≤ A2 →
blocking_bound A1 ≥ blocking_bound A2.
Proof.
move⇒ A1 A2 LEQ.
rewrite /blocking_bound.
apply: bigmax_subset ⇒ tsk_o IN /andP[/andP[OTHER LT] ARR].
by repeat (apply /andP; split) ⇒ //; lia.
Qed.
To use the refined search space with the abstract theorem, we must show
that it still includes all relevant points. To this end, we first observe
that a step in the blocking bound implies the existence of a task that
could release a job with an absolute deadline equal to the absolute
deadline of the job under analysis.
Lemma task_with_equal_deadline_exists :
∀ {A},
priority_inversion_changes_at blocking_bound A →
∃ tsk_o, (tsk_o \in ts)
&& (blocking_relevant tsk_o)
&& (tsk_o != tsk)
&& (D tsk_o == D tsk + A).
Proof.
move⇒ A. rewrite /priority_inversion_changes_at ⇒ NEQ.
have LEQ: blocking_bound A ≤ blocking_bound (A - ε) by apply: blocking_bound_decreasing; lia.
have LT: blocking_bound A < blocking_bound (A - ε) by lia.
move: LT; rewrite /blocking_bound ⇒ LT {LEQ} {NEQ}.
move: (bigmax_witness_diff LT) ⇒ [tsk_o [IN [NOT HOLDS]]].
move: HOLDS ⇒ /andP[REL LTeps].
∃ tsk_o; repeat (apply /andP; split) ⇒ //;
first by apply /eqP ⇒ EQ; move: LTeps; rewrite EQ; lia.
move: NOT; rewrite negb_and ⇒ /orP[/negP // |].
by move: LTeps; rewrite /ε ⇒ LTeps; lia.
Qed.
∀ {A},
priority_inversion_changes_at blocking_bound A →
∃ tsk_o, (tsk_o \in ts)
&& (blocking_relevant tsk_o)
&& (tsk_o != tsk)
&& (D tsk_o == D tsk + A).
Proof.
move⇒ A. rewrite /priority_inversion_changes_at ⇒ NEQ.
have LEQ: blocking_bound A ≤ blocking_bound (A - ε) by apply: blocking_bound_decreasing; lia.
have LT: blocking_bound A < blocking_bound (A - ε) by lia.
move: LT; rewrite /blocking_bound ⇒ LT {LEQ} {NEQ}.
move: (bigmax_witness_diff LT) ⇒ [tsk_o [IN [NOT HOLDS]]].
move: HOLDS ⇒ /andP[REL LTeps].
∃ tsk_o; repeat (apply /andP; split) ⇒ //;
first by apply /eqP ⇒ EQ; move: LTeps; rewrite EQ; lia.
move: NOT; rewrite negb_and ⇒ /orP[/negP // |].
by move: LTeps; rewrite /ε ⇒ LTeps; lia.
Qed.
With the above setup in place, we can show that the search space defined
above by is_in_search_space covers the the more abstract search space
defined by bounded_pi.is_in_search_space.
Lemma search_space_inclusion :
∀ {A L},
bounded_pi.is_in_search_space ts tsk blocking_bound L A →
is_in_search_space L A.
Proof.
move⇒ A L /andP[BOUND STEP].
apply /andP; split ⇒ //; apply /orP.
move: STEP ⇒ /orP[/orP[STEP|RBF] | IBF]; [right| by left| by right].
move: (task_with_equal_deadline_exists STEP) ⇒ [tsk_o /andP[/andP[/andP[IN REL] OTHER] EQ]].
rewrite /bound_on_total_hep_workload_changes_at.
apply /hasP; ∃ tsk_o ⇒ //.
apply /andP; split; first by rewrite eq_sym.
move: EQ. rewrite /D ⇒ /eqP EQ.
rewrite /task_request_bound_function EQ.
move: REL; rewrite /blocking_relevant ⇒ /andP [ARRIVES COST].
rewrite eqn_pmul2l //.
have → : A + task_deadline tsk - (task_deadline tsk + A)
= 0 by lia.
have → : A + ε + task_deadline tsk - (task_deadline tsk + A)
= ε by lia.
by move: (H_valid_arrival_curve tsk_o IN) ⇒ [-> _]; lia.
Qed.
∀ {A L},
bounded_pi.is_in_search_space ts tsk blocking_bound L A →
is_in_search_space L A.
Proof.
move⇒ A L /andP[BOUND STEP].
apply /andP; split ⇒ //; apply /orP.
move: STEP ⇒ /orP[/orP[STEP|RBF] | IBF]; [right| by left| by right].
move: (task_with_equal_deadline_exists STEP) ⇒ [tsk_o /andP[/andP[/andP[IN REL] OTHER] EQ]].
rewrite /bound_on_total_hep_workload_changes_at.
apply /hasP; ∃ tsk_o ⇒ //.
apply /andP; split; first by rewrite eq_sym.
move: EQ. rewrite /D ⇒ /eqP EQ.
rewrite /task_request_bound_function EQ.
move: REL; rewrite /blocking_relevant ⇒ /andP [ARRIVES COST].
rewrite eqn_pmul2l //.
have → : A + task_deadline tsk - (task_deadline tsk + A)
= 0 by lia.
have → : A + ε + task_deadline tsk - (task_deadline tsk + A)
= ε by lia.
by move: (H_valid_arrival_curve tsk_o IN) ⇒ [-> _]; lia.
Qed.
Priority inversion is bounded
In this section, we prove that a priority inversion for task tsk is bounded by the maximum length of non-preemptive segments among the tasks with lower priority.
Since EDF is a JLFP policy, the maximum non-preemptive segment length of any task
that releases a job with an earlier absolute deadline (w.r.t. a given job j) and
non-zero execution cost upper-bounds the maximum possible length of priority
inversion (experienced by said job j).
Using this fact, we prove that the maximum length of a priority inversion of a given
job j is indeed bounded by the defined blocking bound.
Lemma priority_inversion_is_bounded_by_blocking:
∀ j t1 t2,
arrives_in arr_seq j →
job_of_task tsk j →
busy_interval_prefix arr_seq sched j t1 t2 →
max_length_of_priority_inversion j t1 ≤ blocking_bound (job_arrival j - t1).
Proof.
move⇒ j t1 t2 ARR TSK BUSY; rewrite /max_length_of_priority_inversion /blocking_bound.
apply: leq_trans.
exact: priority_inversion_is_bounded_by_max_np_segment.
apply /bigmax_leq_seqP ⇒ j' JINB NOTHEP.
have ARR': arrives_in arr_seq j'
by apply: in_arrivals_implies_arrived; exact: JINB.
apply leq_bigmax_cond_seq with (x := (job_task j')) (F := fun tsk ⇒ task_max_nonpreemptive_segment tsk - 1);
first by apply H_all_jobs_from_taskset.
apply in_arrivals_implies_arrived_between in JINB ⇒ [|//].
move: JINB; move ⇒ /andP [_ TJ'].
repeat (apply/andP; split); last first.
{ rewrite /EDF -ltnNge in NOTHEP.
move: TSK ⇒ /eqP <-.
have ARRLE: job_arrival j' < job_arrival j by apply: (@leq_trans t1).
move: NOTHEP; rewrite /job_deadline /absolute_deadline.job_deadline_from_task_deadline /D.
by lia. }
{ move: NOTHEP ⇒ /andP [_ NZ].
move: (H_valid_job_cost j' ARR'); rewrite /valid_job_cost.
by lia. }
{ apply: non_pathological_max_arrivals; last first.
- exact: ARR'.
- by rewrite /job_of_task.
- by apply H_is_arrival_curve, H_all_jobs_from_taskset, ARR'. }
Qed.
End PriorityInversionIsBounded.
∀ j t1 t2,
arrives_in arr_seq j →
job_of_task tsk j →
busy_interval_prefix arr_seq sched j t1 t2 →
max_length_of_priority_inversion j t1 ≤ blocking_bound (job_arrival j - t1).
Proof.
move⇒ j t1 t2 ARR TSK BUSY; rewrite /max_length_of_priority_inversion /blocking_bound.
apply: leq_trans.
exact: priority_inversion_is_bounded_by_max_np_segment.
apply /bigmax_leq_seqP ⇒ j' JINB NOTHEP.
have ARR': arrives_in arr_seq j'
by apply: in_arrivals_implies_arrived; exact: JINB.
apply leq_bigmax_cond_seq with (x := (job_task j')) (F := fun tsk ⇒ task_max_nonpreemptive_segment tsk - 1);
first by apply H_all_jobs_from_taskset.
apply in_arrivals_implies_arrived_between in JINB ⇒ [|//].
move: JINB; move ⇒ /andP [_ TJ'].
repeat (apply/andP; split); last first.
{ rewrite /EDF -ltnNge in NOTHEP.
move: TSK ⇒ /eqP <-.
have ARRLE: job_arrival j' < job_arrival j by apply: (@leq_trans t1).
move: NOTHEP; rewrite /job_deadline /absolute_deadline.job_deadline_from_task_deadline /D.
by lia. }
{ move: NOTHEP ⇒ /andP [_ NZ].
move: (H_valid_job_cost j' ARR'); rewrite /valid_job_cost.
by lia. }
{ apply: non_pathological_max_arrivals; last first.
- exact: ARR'.
- by rewrite /job_of_task.
- by apply H_is_arrival_curve, H_all_jobs_from_taskset, ARR'. }
Qed.
End PriorityInversionIsBounded.
Response-Time Bound
In this section, we prove that the maximum among the solutions of the response-time bound recurrence is a response-time bound for tsk.
Let L be any positive fixed point of the busy interval recurrence.
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 L A →
∃ (F : duration),
A + F ≥ blocking_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 L A →
∃ (F : duration),
A + F ≥ blocking_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).
Then, using the results for the general RTA for EDF-schedulers, we establish a
response-time bound for the more concrete model of bounded nonpreemptive segments.
Note that in case of the general RTA for EDF-schedulers, we just assume that
the priority inversion is bounded. In this module we provide the preemption model
with bounded nonpreemptive segments and prove that the priority inversion is
bounded.
Theorem uniprocessor_response_time_bound_edf_with_bounded_nonpreemptive_segments:
response_time_bounded_by tsk R.
Proof.
apply: uniprocessor_response_time_bound_edf; eauto 4 with basic_rt_facts.
{ apply: priority_inversion_is_bounded ⇒ //.
apply: priority_inversion_is_bounded_by_blocking. }
- move⇒ A BPI_SP.
by apply H_R_is_maximum, search_space_inclusion.
Qed.
End ResponseTimeBound.
End RTAforEDFwithBoundedNonpreemptiveSegmentsWithArrivalCurves.
response_time_bounded_by tsk R.
Proof.
apply: uniprocessor_response_time_bound_edf; eauto 4 with basic_rt_facts.
{ apply: priority_inversion_is_bounded ⇒ //.
apply: priority_inversion_is_bounded_by_blocking. }
- move⇒ A BPI_SP.
by apply H_R_is_maximum, search_space_inclusion.
Qed.
End ResponseTimeBound.
End RTAforEDFwithBoundedNonpreemptiveSegmentsWithArrivalCurves.