Library prosa.results.fifo.rta
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq path fintype bigop.
Require Import prosa.model.readiness.basic.
Require Import prosa.model.priority.fifo.
Require Import prosa.analysis.facts.priority.fifo.
Require Import prosa.analysis.abstract.ideal_jlfp_rta.
Require Export prosa.analysis.facts.busy_interval.carry_in.
Require Import prosa.model.readiness.basic.
Require Import prosa.model.priority.fifo.
Require Import prosa.analysis.facts.priority.fifo.
Require Import prosa.analysis.abstract.ideal_jlfp_rta.
Require Export prosa.analysis.facts.busy_interval.carry_in.
Abstract RTA for FIFO-schedulers
In this module we instantiate the Abstract Response-Time analysis (aRTA) to FIFO schedulers for real-time tasks with arbitrary arrival models assuming an ideal uni-processor model.
Consider any type of tasks, each characterized by a WCET, a relative
deadline, and a run-to-completion threshold, ...
Context {Task : TaskType}.
Context `{TaskCost Task}.
Context `{TaskDeadline Task}.
Context `{TaskRunToCompletionThreshold Task}.
Context `{TaskCost Task}.
Context `{TaskDeadline Task}.
Context `{TaskRunToCompletionThreshold Task}.
... and any type of jobs associated with these tasks, where each
each job has an arrival time, a cost, and a preemption-point predicate.
Context {Job : JobType}.
Context `{JobTask Job Task}.
Context {Arrival : JobArrival Job}.
Context {Cost : JobCost Job}.
Context `{JobPreemptable Job}.
Context `{JobTask Job Task}.
Context {Arrival : JobArrival Job}.
Context {Cost : JobCost Job}.
Context `{JobPreemptable Job}.
We assume the classic (i.e., Liu & Layland) model of readiness
without jitter or self-suspensions, wherein pending jobs are
always ready.
#[local] Existing Instance basic_ready_instance.
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.
Next, consider any valid ideal uni-processor schedule of this arrival sequence.
Variable sched : schedule (ideal.processor_state Job).
Hypothesis H_valid_schedule : valid_schedule sched arr_seq.
Hypothesis H_valid_schedule : valid_schedule sched arr_seq.
Note that we differentiate between abstract and
classical notions of work-conserving schedules.
Let work_conserving_ab := definitions.work_conserving arr_seq sched.
Let work_conserving_cl := work_conserving.work_conserving arr_seq sched.
Let work_conserving_cl := work_conserving.work_conserving arr_seq sched.
We assume that the schedule is a work-conserving schedule
in the classical sense, and later prove that the hypothesis
about abstract work-conservation also holds.
Assume that a job's cost cannot be larger than its task's WCET.
Consider an arbitrary task set ts.
Next, we assume that all jobs come from the task set.
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) 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.
Consider a valid preemption model...
...and a valid task run-to-completion threshold function. That
is, task_rtct tsk is (1) no larger than tsk's cost, (2) for
any job of task tsk, job_rtct is bounded by task_rtct.
Hypothesis H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold arr_seq tsk.
We also assume that the schedule respects the policy defined by the preemption model.
Hypothesis H_respects_policy_at_preemption_point : respects_JLFP_policy_at_preemption_point arr_seq sched (FIFO Job).
We introduce rbf as an abbreviation of the task request bound function,
which is defined as task_cost(T) × max_arrivals(T,Δ) for a given task T.
For simplicity, let's define some local names.
Let response_time_bounded_by := task_response_time_bound arr_seq sched.
Let number_of_task_arrivals := number_of_task_arrivals arr_seq.
Let number_of_task_arrivals := number_of_task_arrivals arr_seq.
Let L be any positive fixed point of the busy interval recurrence.
Variable L : duration.
Hypothesis H_L_positive : L > 0.
Hypothesis H_fixed_point : L = total_request_bound_function ts L.
Hypothesis H_L_positive : L > 0.
Hypothesis H_fixed_point : L = total_request_bound_function ts L.
To reduce the time complexity of the analysis, we introduce the notion of search space for FIFO.
Intuitively, this corresponds to all "interesting" arrival offsets that the job under
analysis might have with regard to the beginning of its busy-window.
In the case of FIFO, the final search space is the set of offsets less than L
such that there exists a task tsko from ts such that rbf tsko (A) ≠ rbf tsko (A + ε).
Definition is_in_search_space (A : duration) :=
(A < L) && has (fun tsko ⇒ rbf tsko (A) != rbf tsko ( A + ε )) ts.
(A < L) && has (fun tsko ⇒ rbf tsko (A) != rbf tsko ( A + ε )) ts.
Let R be a value that upper-bounds the solution of each
response-time equation, i.e., for any relative arrival time A
in the search space, there exists a corresponding solution F
such that R ≥ F.
Variable R : duration.
Hypothesis H_R_is_maximum:
∀ (A : duration),
is_in_search_space A →
∃ (F : nat),
A + F ≥ \sum_(tsko <- ts) rbf tsko (A + ε) ∧
F ≤ R.
Hypothesis H_R_is_maximum:
∀ (A : duration),
is_in_search_space A →
∃ (F : nat),
A + F ≥ \sum_(tsko <- ts) rbf tsko (A + ε) ∧
F ≤ R.
To use the theorem uniprocessor_response_time_bound from the Abstract RTA module,
we need to specify functions that concretely define the abstract concepts
interference, interfering workload, and IBF.
Instantiation of Interference
We say that job j incurs interference at time t iff it cannot execute due to a higher-or-equal-priority job being scheduled, or if it incurs a priority inversion.Instantiation of Interfering Workload
The interfering workload, in turn, is defined as the sum of the priority inversion function and interfering workload of jobs with higher or equal priority.
Let interfering_workload (j : Job) (t : instant) := ideal_jlfp_rta.interfering_workload arr_seq sched j t.
Finally, we define the interference bound function (IBF). IBF bounds the cumulative
interference incurred by a job. For FIFO, we define IBF as the sum of RBF for all tasks
in the interval A + ε minus the WCET of tsk.
Filling Out Hypotheses Of Abstract RTA Theorem
In this section we prove that all hypotheses necessary to use the abstract theorem are satisfied.
In this section, we prove that, under FIFO scheduling, the cumulative priority inversion experienced
by a job j in any interval within its busy window is always 0. We later use this fact to prove the bound on
the cumulative interference.
Variable j : Job.
Hypothesis H_j_arrives : arrives_in arr_seq j.
Hypothesis H_job_of_tsk : job_task j = tsk.
Hypothesis H_j_arrives : arrives_in arr_seq j.
Hypothesis H_job_of_tsk : job_task j = tsk.
Assume that the job has a positive cost.
Assume the busy interval of the job is given by
[t1,t2)
.
Variable t1 t2 : duration.
Hypothesis H_busy_interval :
definitions.busy_interval sched interference interfering_workload j t1 t2.
Hypothesis H_busy_interval :
definitions.busy_interval sched interference interfering_workload j t1 t2.
Consider any interval
[t1,t1 + Δ)
in the busy interval of j.
We prove that the cumulative priority inversion in the interval
[t1, t1 + Δ)
is 0.
Lemma cumulative_priority_inversion_is_bounded:
cumulative_priority_inversion arr_seq sched j t1 (t1 + Δ) = 0.
End PriorityInversion.
cumulative_priority_inversion arr_seq sched j t1 (t1 + Δ) = 0.
End PriorityInversion.
Using the above lemma, we prove that IBF is indeed an interference bound.
Lemma instantiated_interference_is_bounded :
job_interference_is_bounded_by arr_seq sched tsk interference interfering_workload IBF.
job_interference_is_bounded_by arr_seq sched tsk interference interfering_workload IBF.
Finally, we show that there exists a solution for the response-time equation.
Consider any job j of tsk.
Variable j : Job.
Hypothesis H_j_arrives : arrives_in arr_seq j.
Hypothesis H_job_of_tsk : job_of_task tsk j.
Hypothesis H_positive_cost : 0 < task_cost tsk.
Hypothesis H_j_arrives : arrives_in arr_seq j.
Hypothesis H_job_of_tsk : job_of_task tsk j.
Hypothesis H_positive_cost : 0 < task_cost tsk.
Next, consider any A from the search space (in the abstract sense).
Variable A : nat.
Hypothesis H_A_is_in_abstract_search_space:
search_space.is_in_search_space tsk L IBF A.
Hypothesis H_A_is_in_abstract_search_space:
search_space.is_in_search_space tsk L IBF A.
We prove that A is also in the concrete search space. In other words,
we prove that the abstract search space is a subset of the concrete search space.
Then, there exists a solution for the response-time recurrence (in the abstract sense).
Corollary exists_solution_for_abstract_response_time_recurrence:
∃ (F : nat),
A + F ≥ task_rtct tsk + IBF tsk A (A + F) ∧
F + (task_cost tsk - task_rtct tsk) ≤ R.
End SolutionOfResponseTimeReccurenceExists.
End FillingOutHypothesesOfAbstractRTATheorem.
∃ (F : nat),
A + F ≥ task_rtct tsk + IBF tsk A (A + F) ∧
F + (task_cost tsk - task_rtct tsk) ≤ R.
End SolutionOfResponseTimeReccurenceExists.
End FillingOutHypothesesOfAbstractRTATheorem.