Library prosa.results.prm.fifo.bounded_nps
Require Import prosa.analysis.facts.readiness.basic.
Require Export prosa.model.composite.valid_task_arrival_sequence.
Require Export prosa.analysis.facts.preemption.rtc_threshold.limited.
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.fifo.
Require Export prosa.analysis.abstract.restricted_supply.search_space.fifo_fixpoint.
Require Export prosa.analysis.facts.priority.fifo.
Require Export prosa.analysis.facts.priority.fifo_ahep_bound.
Require Export prosa.analysis.facts.model.sbf.periodic.
Require Export prosa.model.composite.valid_task_arrival_sequence.
Require Export prosa.analysis.facts.preemption.rtc_threshold.limited.
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.fifo.
Require Export prosa.analysis.abstract.restricted_supply.search_space.fifo_fixpoint.
Require Export prosa.analysis.facts.priority.fifo.
Require Export prosa.analysis.facts.priority.fifo_ahep_bound.
Require Export prosa.analysis.facts.model.sbf.periodic.
RTA for FIFO 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
Context {Task : TaskType} `{TaskCost Task} `{MaxArrivals Task}
`{TaskRunToCompletionThreshold Task}.
`{TaskRunToCompletionThreshold Task}.
... and their associated jobs, where each job has a corresponding task
job_task, an execution time job_cost, an arrival time job_arrival,
and a list of preemption points job_preemptable.
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.
We assume that tsk is described by a valid task run-to-completion
threshold. That is, there exists a task parameter task_rtct such that
task_rtct tsk is
Hypothesis H_valid_run_to_completion_threshold :
valid_task_run_to_completion_threshold arr_seq tsk.
valid_task_run_to_completion_threshold arr_seq tsk.
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 complies with some valid preemption model.
Under FIFO scheduling, the specific choice of preemption model does not
matter, since the resulting schedule behaves as if it were non-preemptive:
jobs are executed in arrival order without preemption.
We assume that the schedule respects the FIFO 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 L A →
∃ (F : duration),
prm_sbf Π γ F ≥ total_request_bound_function ts (A + ε)
∧ A + R ≥ F.
∀ (A : duration),
is_in_search_space ts L A →
∃ (F : duration),
prm_sbf Π γ F ≥ total_request_bound_function ts (A + ε)
∧ A + R ≥ F.
Finally, using the abstract restricted-supply analysis, we establish that
any R that satisfies the stated equation is a sound response-time bound
for FIFO scheduling on a unit-speed uniprocessor under the periodic
resource model.
Theorem uniprocessor_response_time_bound_fifo :
∀ (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 [ZERO|POS] := posnP (job_cost js); first by rewrite /job_response_time_bound /completed_by ZERO.
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 POStsk: 0 < task_cost tsk
by move: TSKs ⇒ /eqP <-; apply: leq_trans; [apply POS | apply H_valid_task_arrival_sequence].
eapply uniprocessor_response_time_bound_restricted_supply
with (L := L) (intra_IBF := fun A Δ ⇒ (\sum_(tsko <- ts) task_request_bound_function tsko (A + ε)) - task_cost tsk) ⇒ //.
- exact: instantiated_i_and_w_are_coherent_with_schedule.
- eapply busy_intervals_are_bounded_rs_jlfp with (blocking_bound := fun _ ⇒ 0)=> //.
+ exact: instantiated_i_and_w_are_coherent_with_schedule.
+ by apply: FIFO_implies_no_service_inversion.
- apply: valid_pred_sbf_switch_predicate; last first.
+ by apply prm_sbf_valid.
+ by move ⇒ ? ? ? ? [? ?]; split ⇒ //.
- move ⇒ t1 t2 Δ j ARR TSK BUSY LT NCOMPL A OFF.
move: (OFF _ _ BUSY) ⇒ EQA; subst A.
have [ZERO|POSj] := posnP (job_cost j).
{ by exfalso; rewrite /completed_by ZERO in NCOMPL. }
eapply leq_trans; first by eapply cumulative_intra_interference_split ⇒ //.
rewrite -[leqRHS]add0n leq_add //.
+ rewrite (leqRW (service_inversion_widen _ _ _ _ _ _ _ _ _)).
× apply: FIFO_implies_no_service_inversion ⇒ //.
by apply instantiated_busy_interval_equivalent_busy_interval ⇒ //.
× by done.
× by done.
+ unshelve apply: bound_on_hep_workload; (try apply H_fixed_point).
all: try apply H_L_positive.
all: try done.
apply instantiated_busy_interval_equivalent_busy_interval ⇒ //.
- apply: soln_abstract_response_time_recurrence ⇒ //.
+ by exact: prm_sbf_monotone.
+ by apply: non_pathological_max_arrivals =>//; apply H_valid_task_arrival_sequence.
Qed.
End RTAforFIFOModelwithArrivalCurves.
∀ (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 [ZERO|POS] := posnP (job_cost js); first by rewrite /job_response_time_bound /completed_by ZERO.
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 POStsk: 0 < task_cost tsk
by move: TSKs ⇒ /eqP <-; apply: leq_trans; [apply POS | apply H_valid_task_arrival_sequence].
eapply uniprocessor_response_time_bound_restricted_supply
with (L := L) (intra_IBF := fun A Δ ⇒ (\sum_(tsko <- ts) task_request_bound_function tsko (A + ε)) - task_cost tsk) ⇒ //.
- exact: instantiated_i_and_w_are_coherent_with_schedule.
- eapply busy_intervals_are_bounded_rs_jlfp with (blocking_bound := fun _ ⇒ 0)=> //.
+ exact: instantiated_i_and_w_are_coherent_with_schedule.
+ by apply: FIFO_implies_no_service_inversion.
- apply: valid_pred_sbf_switch_predicate; last first.
+ by apply prm_sbf_valid.
+ by move ⇒ ? ? ? ? [? ?]; split ⇒ //.
- move ⇒ t1 t2 Δ j ARR TSK BUSY LT NCOMPL A OFF.
move: (OFF _ _ BUSY) ⇒ EQA; subst A.
have [ZERO|POSj] := posnP (job_cost j).
{ by exfalso; rewrite /completed_by ZERO in NCOMPL. }
eapply leq_trans; first by eapply cumulative_intra_interference_split ⇒ //.
rewrite -[leqRHS]add0n leq_add //.
+ rewrite (leqRW (service_inversion_widen _ _ _ _ _ _ _ _ _)).
× apply: FIFO_implies_no_service_inversion ⇒ //.
by apply instantiated_busy_interval_equivalent_busy_interval ⇒ //.
× by done.
× by done.
+ unshelve apply: bound_on_hep_workload; (try apply H_fixed_point).
all: try apply H_L_positive.
all: try done.
apply instantiated_busy_interval_equivalent_busy_interval ⇒ //.
- apply: soln_abstract_response_time_recurrence ⇒ //.
+ by exact: prm_sbf_monotone.
+ by apply: non_pathological_max_arrivals =>//; apply H_valid_task_arrival_sequence.
Qed.
End RTAforFIFOModelwithArrivalCurves.