Library prosa.results.edf.rta.limited_preemptive
Require Export prosa.results.edf.rta.bounded_nps.
Require Export prosa.analysis.facts.preemption.rtc_threshold.limited.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq path fintype bigop.
Require Export prosa.analysis.facts.preemption.rtc_threshold.limited.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq path fintype bigop.
RTA for EDF with Fixed Preemption Points
In this module we prove the RTA theorem for EDF-schedulers with fixed preemption points.
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 task model with fixed preemption points.
Require Import prosa.model.preemption.limited_preemptive.
Require Import prosa.model.task.preemption.limited_preemptive.
Require Import prosa.model.task.preemption.limited_preemptive.
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.
Next, we assume we have the model with fixed preemption points.
I.e., each task is divided into a number of non-preemptive segments
by inserting statically predefined preemption points.
Context `{JobPreemptionPoints Job}.
Context `{TaskPreemptionPoints Task}.
Hypothesis H_valid_model_with_fixed_preemption_points:
valid_fixed_preemption_points_model arr_seq ts.
Context `{TaskPreemptionPoints Task}.
Hypothesis H_valid_model_with_fixed_preemption_points:
valid_fixed_preemption_points_model arr_seq ts.
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 ideal uni-processor schedule with limited
preemptions of this arrival sequence ...
Variable sched : schedule (ideal.processor_state Job).
Hypothesis H_jobs_come_from_arrival_sequence:
jobs_come_from_arrival_sequence sched arr_seq.
Hypothesis H_schedule_with_limited_preemptions:
schedule_respects_preemption_model arr_seq sched.
Hypothesis H_jobs_come_from_arrival_sequence:
jobs_come_from_arrival_sequence sched arr_seq.
Hypothesis H_schedule_with_limited_preemptions:
schedule_respects_preemption_model arr_seq sched.
... where jobs do not execute before their arrival or 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 we have sequential tasks, i.e, jobs from the
same task execute in the order of their arrival.
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).
We define a bound for the priority inversion caused by jobs with lower priority.
Let blocking_bound :=
\max_(tsk_other <- ts | (tsk_other != tsk) && (task_deadline tsk_other > task_deadline tsk))
(task_max_nonpreemptive_segment tsk_other - ε).
\max_(tsk_other <- ts | (tsk_other != tsk) && (task_deadline tsk_other > task_deadline tsk))
(task_max_nonpreemptive_segment tsk_other - ε).
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
Let task_rbf_changes_at A := task_rbf_changes_at tsk A.
Let bound_on_total_hep_workload_changes_at :=
bound_on_total_hep_workload_changes_at ts tsk.
Let is_in_search_space A :=
(A < L) && (task_rbf_changes_at A || bound_on_total_hep_workload_changes_at A).
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 = blocking_bound
+ (task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε))
+ bound_on_total_hep_workload A (A + F) ∧
F + (task_last_nonpr_segment tsk - ε) ≤ R.
Hypothesis H_R_is_maximum:
∀ (A : duration),
is_in_search_space A →
∃ (F : duration),
A + F = blocking_bound
+ (task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε))
+ bound_on_total_hep_workload A (A + F) ∧
F + (task_last_nonpr_segment tsk - ε) ≤ R.
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 fixed preemption points.
Let response_time_bounded_by := task_response_time_bound arr_seq sched.
Theorem uniprocessor_response_time_bound_edf_with_fixed_preemption_points:
response_time_bounded_by tsk R.
Proof.
move: (H_valid_model_with_fixed_preemption_points) ⇒ [MLP [BEG [END [INCR [HYP1 [HYP2 HYP3]]]]]].
move: (MLP) ⇒ [BEGj [ENDj _]].
case: (posnP (task_cost tsk)) ⇒ [ZERO|POSt].
{ intros j ARR TSK.
move: (H_valid_job_cost _ ARR) ⇒ POSt.
move: POSt; rewrite /valid_job_cost TSK ZERO leqn0; move ⇒ /eqP Z.
by rewrite /job_response_time_bound /completed_by Z.
}
eapply uniprocessor_response_time_bound_edf_with_bounded_nonpreemptive_segments with (L0 := L).
all: eauto 2 with basic_facts.
{ rewrite subKn; first by done.
rewrite /task_last_nonpr_segment -(leq_add2r 1) subn1 !addn1 prednK; last first.
{ rewrite /last0 -nth_last.
apply HYP3; try by done.
rewrite -(ltn_add2r 1) !addn1 prednK //.
move: (number_of_preemption_points_in_task_at_least_two
_ _ H_valid_model_with_fixed_preemption_points _ H_tsk_in_ts POSt) ⇒ Fact2.
move: (Fact2) ⇒ Fact3.
by rewrite size_of_seq_of_distances // addn1 ltnS // in Fact2.
}
{ apply leq_trans with (task_max_nonpreemptive_segment tsk).
- by apply last_of_seq_le_max_of_seq.
- rewrite -END; last by done.
apply ltnW; rewrite ltnS; try done.
by apply max_distance_in_seq_le_last_element_of_seq; eauto 2.
}
}
Qed.
End RTAforFixedPreemptionPointsModelwithArrivalCurves.