Library prosa.results.prm.edf.fully_preemptive
Require Import prosa.analysis.facts.readiness.basic.
Require Export prosa.model.composite.valid_task_arrival_sequence.
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.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.sbf.periodic.
Require Export prosa.model.composite.valid_task_arrival_sequence.
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.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.sbf.periodic.
RTA for Fully Preemptive EDF Scheduling on Uniprocessors under the Periodic Resource Model
Defining the System Model
- tasks, jobs, and their parameters,
- the task set and the task under analysis,
- the processor model,
- the sequence of job arrivals,
- the absence of self-suspensions,
- an arbitrary schedule of the task set, and finally,
- the resource-supply model.
Tasks and Jobs
... and their associated jobs, where each job has a corresponding task
job_task, an execution time job_cost, and an arrival time
job_arrival.
We 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.
Context `{PState : ProcessorState Job}.
Hypothesis H_uniprocessor_proc_model : uniprocessor_model PState.
Hypothesis H_unit_supply_proc_model : unit_supply_proc_model PState.
Hypothesis H_consumed_supply_proc_model : fully_consuming_proc_model PState.
Hypothesis H_uniprocessor_proc_model : uniprocessor_model PState.
Hypothesis H_unit_supply_proc_model : unit_supply_proc_model PState.
Hypothesis H_consumed_supply_proc_model : fully_consuming_proc_model PState.
The Job Arrival Sequence
Variable arr_seq : arrival_sequence Job.
Hypothesis H_valid_task_arrival_sequence : valid_task_arrival_sequence ts arr_seq.
Hypothesis H_valid_task_arrival_sequence : valid_task_arrival_sequence ts arr_seq.
Absence of Self-Suspensions
The Schedule
Variable sched : schedule PState.
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.
We assume that the schedule respects the given EDF scheduling policy.
Periodic Resource Model
[Π⋅k, Π⋅(k+1))
, the processor
provides at least γ units of supply. Furthermore, let prm_sbf Π γ
denote the corresponding SBF defined in the paper, which, as proven in the
same paper and verified in prosa.analysis.facts.model.sbf.periodic, is a
valid SBF.
Maximum Length of a Busy Interval
Definition busy_window_recurrence_solution (L : duration) :=
L > 0
∧ prm_sbf Π γ L ≥ total_request_bound_function ts L.
L > 0
∧ prm_sbf Π γ L ≥ total_request_bound_function ts L.
Response-Time Bound
Definition rta_recurrence_solution L R :=
∀ (A : duration),
is_in_search_space ts tsk L A →
∃ (F : duration),
prm_sbf Π γ F ≥ task_request_bound_function tsk (A + ε)
+ bound_on_athep_workload ts tsk A F
∧ A + R ≥ F.
∀ (A : duration),
is_in_search_space ts tsk L A →
∃ (F : duration),
prm_sbf Π γ F ≥ task_request_bound_function tsk (A + ε)
+ bound_on_athep_workload ts tsk A F
∧ A + R ≥ F.
Finally, using the sequential variant of abstract restricted-supply
analysis, we establish that, given a bound on the maximum busy-window
length L, any such R is indeed a sound response-time bound for task
tsk under fully-preemptive EDF scheduling on a unit-speed uniprocessor
under the periodic resource model.
Theorem uniprocessor_response_time_bound_fully_preemptive_edf :
∀ (L : duration),
busy_window_recurrence_solution L →
∀ (R : duration),
rta_recurrence_solution L R →
task_response_time_bound arr_seq sched tsk R.
Proof.
move⇒ L [BW_POS BW_SOL] R SOL js ARRs TSKs.
have VBSBF : valid_busy_sbf arr_seq sched tsk (prm_sbf Π γ).
{ by apply: valid_pred_sbf_switch_predicate; last apply prm_sbf_valid. }
have US : unit_supply_bound_function (prm_sbf Π γ) by exact: prm_sbf_unit.
have [ZERO|POS] := posnP (job_cost js); first by rewrite /job_response_time_bound /completed_by ZERO.
have POStsk: 0 < task_cost tsk
by move: TSKs ⇒ /eqP <-; apply: leq_trans; [apply POS | apply H_valid_task_arrival_sequence].
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.
+ exact: instantiated_i_and_w_are_coherent_with_schedule.
+ apply: service_inversion_is_bounded ⇒ // ⇒ ? ? ? ? ? ?.
exact: nonpreemptive_segments_bounded_by_blocking.
+ move: BW_SOL; rewrite BLOCK add0n; lia.
- apply: valid_pred_sbf_switch_predicate; last first.
+ by apply prm_sbf_valid.
+ by move ⇒ ? ? ? ? [? ?]; split ⇒ //.
- 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: (SOL A) ⇒ [].
+ apply: search_space_sub ⇒ //=.
by apply: non_pathological_max_arrivals =>//; apply H_valid_task_arrival_sequence.
+ move ⇒ F [FIX LE]; ∃ F; split ⇒ //.
rewrite /task_intra_IBF /task_rtct /fully_preemptive_rtc_threshold.
rewrite BLOCK subnn //= add0n addn0 subn0; split.
× by move: FIX; lia.
× exact: prm_sbf_monotone.
Qed.
End RTAforFullyPreemptiveEDFModelwithArrivalCurves.
∀ (L : duration),
busy_window_recurrence_solution L →
∀ (R : duration),
rta_recurrence_solution L R →
task_response_time_bound arr_seq sched tsk R.
Proof.
move⇒ L [BW_POS BW_SOL] R SOL js ARRs TSKs.
have VBSBF : valid_busy_sbf arr_seq sched tsk (prm_sbf Π γ).
{ by apply: valid_pred_sbf_switch_predicate; last apply prm_sbf_valid. }
have US : unit_supply_bound_function (prm_sbf Π γ) by exact: prm_sbf_unit.
have [ZERO|POS] := posnP (job_cost js); first by rewrite /job_response_time_bound /completed_by ZERO.
have POStsk: 0 < task_cost tsk
by move: TSKs ⇒ /eqP <-; apply: leq_trans; [apply POS | apply H_valid_task_arrival_sequence].
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.
+ exact: instantiated_i_and_w_are_coherent_with_schedule.
+ apply: service_inversion_is_bounded ⇒ // ⇒ ? ? ? ? ? ?.
exact: nonpreemptive_segments_bounded_by_blocking.
+ move: BW_SOL; rewrite BLOCK add0n; lia.
- apply: valid_pred_sbf_switch_predicate; last first.
+ by apply prm_sbf_valid.
+ by move ⇒ ? ? ? ? [? ?]; split ⇒ //.
- 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: (SOL A) ⇒ [].
+ apply: search_space_sub ⇒ //=.
by apply: non_pathological_max_arrivals =>//; apply H_valid_task_arrival_sequence.
+ move ⇒ F [FIX LE]; ∃ F; split ⇒ //.
rewrite /task_intra_IBF /task_rtct /fully_preemptive_rtc_threshold.
rewrite BLOCK subnn //= add0n addn0 subn0; split.
× by move: FIX; lia.
× exact: prm_sbf_monotone.
Qed.
End RTAforFullyPreemptiveEDFModelwithArrivalCurves.