Library prosa.results.rs.edf.fully_preemptive
Require Import prosa.analysis.facts.readiness.basic.
Require Export prosa.analysis.facts.model.restricted_supply.schedule.
Require Export prosa.analysis.facts.preemption.task.preemptive.
Require Export prosa.analysis.facts.preemption.rtc_threshold.preemptive.
Require Export prosa.analysis.abstract.restricted_supply.task_intra_interference_bound.
Require Export prosa.analysis.abstract.restricted_supply.bounded_bi.jlfp.
Require Export prosa.analysis.abstract.restricted_supply.search_space.edf.
Require Export prosa.analysis.facts.model.task_cost.
Require Export prosa.analysis.facts.priority.edf.
Require Export prosa.analysis.facts.blocking_bound.edf.
Require Export prosa.analysis.facts.workload.edf_athep_bound.
Require Export prosa.analysis.facts.model.restricted_supply.schedule.
Require Export prosa.analysis.facts.preemption.task.preemptive.
Require Export prosa.analysis.facts.preemption.rtc_threshold.preemptive.
Require Export prosa.analysis.abstract.restricted_supply.task_intra_interference_bound.
Require Export prosa.analysis.abstract.restricted_supply.bounded_bi.jlfp.
Require Export prosa.analysis.abstract.restricted_supply.search_space.edf.
Require Export prosa.analysis.facts.model.task_cost.
Require Export prosa.analysis.facts.priority.edf.
Require Export prosa.analysis.facts.blocking_bound.edf.
Require Export prosa.analysis.facts.workload.edf_athep_bound.
RTA for Fully Preemptive EDF Scheduling on Restricted-Supply Uniprocessors
Defining the System Model
- processor model,
- tasks, jobs, and their parameters,
- the sequence of job arrivals,
- worst-case execution time (WCET) and the absence of self-suspensions,
- the set of tasks under analysis,
- the task under analysis, and, finally,
- an arbitrary schedule of the task set.
Processor Model
... where the minimum amount of supply is lower-bounded via a
monotone unit-supply-bound function SBF.
Context {SBF : SupplyBoundFunction}.
Hypothesis H_SBF_monotone : sbf_is_monotone SBF.
Hypothesis H_unit_SBF : unit_supply_bound_function SBF.
Hypothesis H_SBF_monotone : sbf_is_monotone SBF.
Hypothesis H_unit_SBF : unit_supply_bound_function SBF.
Tasks and Jobs
Context {Task : TaskType}.
Context `{TaskCost Task}.
Context `{TaskDeadline Task}.
Context `{MaxArrivals Task}.
Context `{TaskCost Task}.
Context `{TaskDeadline Task}.
Context `{MaxArrivals Task}.
... and any type of jobs associated with these tasks, where each
job has a task job_task, a cost job_cost, and an arrival
time job_arrival.
Context {Job : JobType}.
Context `{JobTask Job Task}.
Context `{JobCost Job}.
Context `{JobArrival Job}.
Context `{JobTask Job Task}.
Context `{JobCost Job}.
Context `{JobArrival Job}.
Furthermore, assume that jobs and tasks are fully preemptive.
#[local] Existing Instance fully_preemptive_job_model.
#[local] Existing Instance fully_preemptive_task_model.
#[local] Existing Instance fully_preemptive_rtc_threshold.
#[local] Existing Instance fully_preemptive_task_model.
#[local] Existing Instance fully_preemptive_rtc_threshold.
The Job Arrival Sequence
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.
Absence of Self-Suspensions and WCET Compliance
We further require that a job's cost cannot exceed its task's stated WCET.
... and assume that all jobs stem from tasks in this task set.
We assume that max_arrivals is a family of valid arrival
curves that constrains the arrival sequence arr_seq, i.e., for
any task tsk in ts, max_arrival tsk is (1) an arrival
bound of tsk, and ...
... (2) a monotonic function that equals 0 for the empty interval delta = 0.
The Schedule
Variable sched : schedule (rs_processor_state Job).
Hypothesis H_valid_schedule : valid_schedule sched arr_seq.
Hypothesis H_work_conserving : work_conserving arr_seq sched.
Hypothesis H_valid_schedule : valid_schedule sched arr_seq.
Hypothesis H_work_conserving : work_conserving arr_seq sched.
... and assume that the schedule respects the EDF policy.
Last but not least, we assume that SBF properly characterizes
all busy intervals (w.r.t. task tsk) in sched. That is, (1)
SBF 0 = 0 and (2) for any duration Δ, at least SBF Δ
supply is available in any busy-interval prefix of length
Δ.
Length of Busy Interval
Variable L : duration.
Hypothesis H_L_positive : 0 < L.
Hypothesis H_fixed_point : total_request_bound_function ts L ≤ SBF L.
Hypothesis H_L_positive : 0 < L.
Hypothesis H_fixed_point : total_request_bound_function ts L ≤ SBF L.
Response-Time Bound
Variable R : duration.
Hypothesis H_R_is_maximum :
∀ (A : duration),
is_in_search_space ts tsk L A →
∃ (F : duration),
A ≤ F ≤ A + R
∧ task_request_bound_function tsk (A + ε) + bound_on_athep_workload ts tsk A F ≤ SBF F.
Hypothesis H_R_is_maximum :
∀ (A : duration),
is_in_search_space ts tsk L A →
∃ (F : duration),
A ≤ F ≤ A + R
∧ task_request_bound_function tsk (A + ε) + bound_on_athep_workload ts tsk A F ≤ SBF F.
Finally, using the sequential variant of abstract
restricted-supply analysis, we establish that any such R is a
sound response-time bound for the concrete model of
fully-preemptive EDF scheduling with arbitrary supply
restrictions.
Theorem uniprocessor_response_time_bound_fully_preemptive_edf :
task_response_time_bound arr_seq sched tsk R.
Proof.
move⇒ js ARRs TSKs.
have [ZERO|POS] := posnP (job_cost js);
first by rewrite /job_response_time_bound /completed_by ZERO.
have READ : work_bearing_readiness arr_seq sched by done.
have BLOCK: ∀ tsk A, blocking_bound ts tsk A = 0.
{ by move⇒ A tsk2; rewrite /blocking_bound /parameters.task_max_nonpreemptive_segment
/fully_preemptive_task_model subnn big1_eq. }
eapply uniprocessor_response_time_bound_restricted_supply_seq with (L := L) ⇒ //.
- exact: instantiated_i_and_w_are_coherent_with_schedule.
- exact: EDF_implies_sequential_tasks.
- exact: instantiated_interference_and_workload_consistent_with_sequential_tasks.
- eapply busy_intervals_are_bounded_rs_jlfp; try done.
+ apply: service_inversion_is_bounded ⇒ // ⇒ ? ? ? ? ? ?.
exact: nonpreemptive_segments_bounded_by_blocking.
+ by rewrite BLOCK add0n; apply H_fixed_point.
- apply: valid_pred_sbf_switch_predicate; last by exact: H_valid_SBF.
move ⇒ ? ? ? ? [? ?]; split ⇒ //.
by apply instantiated_busy_interval_prefix_equivalent_busy_interval_prefix.
- apply: instantiated_task_intra_interference_is_bounded; eauto 1 ⇒ //; first last.
+ by (apply: bound_on_athep_workload_is_valid; try apply H_fixed_point) ⇒ //.
+ apply: service_inversion_is_bounded ⇒ // ⇒ jo t1 t2 ARRo TSKo BUSYo.
by apply: nonpreemptive_segments_bounded_by_blocking ⇒ //.
- move ⇒ A SP.
move: (H_R_is_maximum A) ⇒ [].
+ by apply: search_space_sub ⇒ //.
+ move ⇒ F [/andP [_ LE] FIX]; ∃ F; split ⇒ //.
rewrite /task_intra_IBF /task_rtct /fully_preemptive_rtc_threshold.
by rewrite BLOCK subnn //= add0n addn0 subn0.
Qed.
End RTAforFullyPreemptiveEDFModelwithArrivalCurves.
task_response_time_bound arr_seq sched tsk R.
Proof.
move⇒ js ARRs TSKs.
have [ZERO|POS] := posnP (job_cost js);
first by rewrite /job_response_time_bound /completed_by ZERO.
have READ : work_bearing_readiness arr_seq sched by done.
have BLOCK: ∀ tsk A, blocking_bound ts tsk A = 0.
{ by move⇒ A tsk2; rewrite /blocking_bound /parameters.task_max_nonpreemptive_segment
/fully_preemptive_task_model subnn big1_eq. }
eapply uniprocessor_response_time_bound_restricted_supply_seq with (L := L) ⇒ //.
- exact: instantiated_i_and_w_are_coherent_with_schedule.
- exact: EDF_implies_sequential_tasks.
- exact: instantiated_interference_and_workload_consistent_with_sequential_tasks.
- eapply busy_intervals_are_bounded_rs_jlfp; try done.
+ apply: service_inversion_is_bounded ⇒ // ⇒ ? ? ? ? ? ?.
exact: nonpreemptive_segments_bounded_by_blocking.
+ by rewrite BLOCK add0n; apply H_fixed_point.
- apply: valid_pred_sbf_switch_predicate; last by exact: H_valid_SBF.
move ⇒ ? ? ? ? [? ?]; split ⇒ //.
by apply instantiated_busy_interval_prefix_equivalent_busy_interval_prefix.
- apply: instantiated_task_intra_interference_is_bounded; eauto 1 ⇒ //; first last.
+ by (apply: bound_on_athep_workload_is_valid; try apply H_fixed_point) ⇒ //.
+ apply: service_inversion_is_bounded ⇒ // ⇒ jo t1 t2 ARRo TSKo BUSYo.
by apply: nonpreemptive_segments_bounded_by_blocking ⇒ //.
- move ⇒ A SP.
move: (H_R_is_maximum A) ⇒ [].
+ by apply: search_space_sub ⇒ //.
+ move ⇒ F [/andP [_ LE] FIX]; ∃ F; split ⇒ //.
rewrite /task_intra_IBF /task_rtct /fully_preemptive_rtc_threshold.
by rewrite BLOCK subnn //= add0n addn0 subn0.
Qed.
End RTAforFullyPreemptiveEDFModelwithArrivalCurves.