Skip to content
Snippets Groups Projects
Commit 0d392409 authored by Sergey Bozhko's avatar Sergey Bozhko :eyes:
Browse files

EDF instantiations

Add instantiations of aRTA for
(1) fully preemptive,
(2) fully non-preemptive,
(3) limited preemptions,
(4) and floating non-preemptive regions EDF models
parent 11dba7fb
No related branches found
No related tags found
No related merge requests found
From rt.util Require Import all.
From rt.restructuring.behavior Require Import all.
From rt.restructuring.analysis.basic_facts Require Import all.
From rt.restructuring.model Require Import job task workload processor.ideal readiness.basic.
From rt.restructuring.model.arrival Require Import arrival_curves.
From rt.restructuring.model Require Import preemption.floating.
From rt.restructuring.model.schedule Require Import
work_conserving priority_based.priorities priority_based.edf priority_based.preemption_aware.
From rt.restructuring.analysis.arrival Require Import workload_bound rbf.
From rt.restructuring.analysis.edf.rta Require Import nonpr_reg.response_time_bound.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq path fintype bigop.
(** * RTA for Model with Floating Non-Preemptive Regions *)
(** In this module we prove the RTA theorem for floating non-preemptive regions EDF model. *)
Section RTAforModelWithFloatingNonpreemptiveRegionsWithArrivalCurves.
(** Consider any type of tasks ... *)
Context {Task : TaskType}.
Context `{TaskCost Task}.
Context `{TaskDeadline Task}.
(** ... and any type of jobs associated with these tasks. *)
Context {Job : JobType}.
Context `{JobTask Job Task}.
Context `{JobArrival Job}.
Context `{JobCost Job}.
(** For clarity, let's denote the relative deadline of a task as D. *)
Let D tsk := task_deadline tsk.
(** Consider the EDF policy that indicates a higher-or-equal priority relation. *)
Let EDF := EDF Task 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.
(** Assume we have the model with floating nonpreemptive regions.
I.e., for each task only the length of the maximal nonpreemptive
segment is known _and_ each job level is divided into a number
of nonpreemptive segments by inserting preemption points. *)
Context `{JobPreemptionPoints Job}
`{TaskMaxNonpreemptiveSegment Task}.
Hypothesis H_valid_task_model_with_floating_nonpreemptive_regions:
valid_model_with_floating_nonpreemptive_regions arr_seq.
(** Consider an arbitrary task set ts, ... *)
Variable ts : list Task.
(** ... assume that all jobs come from this task set, ... *)
Hypothesis H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts.
(** ... and the cost of a job cannot be larger than the task cost. *)
Hypothesis H_job_cost_le_task_cost:
cost_of_jobs_from_arrival_sequence_le_task_cost arr_seq.
(** 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.
(** Let tsk be any task in ts that is to be analyzed. *)
Variable tsk : Task.
Hypothesis H_tsk_in_ts : tsk \in ts.
(** Next, consider any ideal uniprocessor 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:
valid_schedule_with_limited_preemptions 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.
(** Assume we have sequential tasks, i.e, jobs from the
same task execute in the order of their arrival. *)
Hypothesis H_sequential_tasks : sequential_tasks sched.
(** Next, we assume that the schedule is a work-conserving schedule... *)
Hypothesis H_work_conserving : work_conserving arr_seq sched.
(** ... and the schedule respects the policy defined by the
job_preemptable function (i.e., jobs have bounded nonpreemptive
segments). *)
Hypothesis H_respects_policy : respects_policy_at_preemption_point arr_seq sched.
(** Let's define some local names for clarity. *)
Let response_time_bounded_by :=
task_response_time_bound arr_seq sched.
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.
(** We introduce the abbreviation "rbf" for the task request bound function,
which is defined as [task_cost(T) × max_arrivals(T,Δ)] for a task T. *)
Let rbf := task_request_bound_function.
(** Next, we introduce task_rbf as an abbreviation
for the task request bound function of task tsk. *)
Let task_rbf := rbf tsk.
(** Using the sum of individual request bound functions, we define the request bound
function of all tasks (total request bound function). *)
Let total_rbf := total_request_bound_function ts.
(** We define a bound for the priority inversion caused by jobs with lower priority. *)
Definition blocking_bound :=
\max_(tsk_other <- ts | (tsk_other != tsk) && (D tsk_other > D 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 + ε) + D tsk - D tsk_o) Δ).
(** 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_rbf L.
(** To reduce the time complexity of the analysis, recall the notion of search space. *)
Let is_in_search_space (A : duration) :=
(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:
forall (A : duration),
is_in_search_space A ->
exists (F : duration),
A + F = blocking_bound + task_rbf (A + ε) + bound_on_total_hep_workload A (A + F) /\
F <= R.
(** Now, we can leverage the results for the abstract model with
bounded nonpreemptive segments to establish a response-time
bound for the more concrete model with floating nonpreemptive
regions. *)
Theorem uniprocessor_response_time_bound_edf_with_floating_nonpreemptive_regions:
response_time_bounded_by tsk R.
Proof.
move: (H_valid_task_model_with_floating_nonpreemptive_regions) => [LIMJ JMLETM].
move: (LIMJ) => [BEG [END _]].
eapply uniprocessor_response_time_bound_edf_with_bounded_nonpreemptive_segments with (L0 := L).
all: eauto 2 with basic_facts.
{ rewrite subnn.
intros A SP.
apply H_R_is_maximum in SP.
move: SP => [F [EQ LE]].
exists F.
by rewrite subn0 addn0; split.
}
Qed.
End RTAforModelWithFloatingNonpreemptiveRegionsWithArrivalCurves.
\ No newline at end of file
From rt.util Require Import all.
From rt.restructuring.behavior Require Import all.
From rt.restructuring.analysis.basic_facts Require Import all.
From rt.restructuring.model Require Import job task workload processor.ideal readiness.basic.
From rt.restructuring.model.arrival Require Import arrival_curves.
From rt.restructuring.model Require Import preemption.limited.
From rt.restructuring.model.schedule Require Import
work_conserving priority_based.priorities priority_based.edf priority_based.preemption_aware.
From rt.restructuring.analysis.arrival Require Import workload_bound rbf.
From rt.restructuring.analysis.edf.rta Require Import nonpr_reg.response_time_bound.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq path fintype bigop.
(** * RTA for EDF-schedulers with Fixed Premption Points *)
(** In this module we prove the RTA theorem for EDF-schedulers with fixed preemption points. *)
Section RTAforFixedPreemptionPointsModelwithArrivalCurves.
(** Consider any type of tasks ... *)
Context {Task : TaskType}.
Context `{TaskCost Task}.
Context `{TaskDeadline Task}.
(** ... and any type of jobs associated with these tasks. *)
Context {Job : JobType}.
Context `{JobTask Job Task}.
Context `{JobArrival Job}.
Context `{JobCost Job}.
(** For clarity, let's denote the relative deadline of a task as D. *)
Let D tsk := task_deadline tsk.
(** Consider the EDF policy that indicates a higher-or-equal priority relation. *)
Let EDF := EDF Task 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.
(** Consider an arbitrary task set ts, ... *)
Variable ts : list Task.
(** ... assume that all jobs come from this task set, ... *)
Hypothesis H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts.
(** ... and the cost of a job cannot be larger than the task cost. *)
Hypothesis H_job_cost_le_task_cost:
cost_of_jobs_from_arrival_sequence_le_task_cost arr_seq.
(** Next, we assume we have the model with fixed preemption points.
I.e., each task is divided into a number of nonpreemptive segments
by inserting staticaly 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.
(** 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.
(** Let tsk be any task in ts that is to be analyzed. *)
Variable tsk : Task.
Hypothesis H_tsk_in_ts : tsk \in ts.
(** Next, consider any ideal uniprocessor schedule with limited
preemptionsof 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:
valid_schedule_with_limited_preemptions 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.
(** Assume we have sequential tasks, i.e, jobs from the
same task execute in the order of their arrival. *)
Hypothesis H_sequential_tasks : sequential_tasks sched.
(** Next, we assume that the schedule is a work-conserving schedule... *)
Hypothesis H_work_conserving : work_conserving arr_seq sched.
(** ... and the schedule respects the policy defined by the
job_preemptable function (i.e., jobs have bounded nonpreemptive
segments). *)
Hypothesis H_respects_policy : respects_policy_at_preemption_point arr_seq sched.
(** Let's define some local names for clarity. *)
Let response_time_bounded_by := task_response_time_bound arr_seq sched.
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.
(** We introduce the abbreviation "rbf" for the task request bound function,
which is defined as [task_cost(T) × max_arrivals(T,Δ)] for a task T. *)
Let rbf := task_request_bound_function.
(** Next, we introduce task_rbf as an abbreviation
for the task request bound function of task tsk. *)
Let task_rbf := rbf tsk.
(** Using the sum of individual request bound functions, we define the request bound
function of all tasks (total request bound function). *)
Let total_rbf := total_request_bound_function ts.
(** We define a bound for the priority inversion caused by jobs with lower priority. *)
Let blocking_bound :=
\max_(tsk_other <- ts | (tsk_other != tsk) && (D tsk_other > D 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 + ε) + D tsk - D tsk_o) Δ).
(** 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_rbf L.
(** To reduce the time complexity of the analysis, recall the notion of search space. *)
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:
forall (A : duration),
is_in_search_space A ->
exists (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 nonpreemptive segments
to establish a response-time bound for the more concrete model of fixed preemption points. *)
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_job_cost_le_task_cost _ ARR) => POSt.
move: POSt; rewrite /job_cost_le_task_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.
\ No newline at end of file
From rt.util Require Import all.
From rt.restructuring.behavior Require Import all.
From rt.restructuring.analysis.basic_facts Require Import all.
From rt.restructuring.model Require Import job task workload processor.ideal readiness.basic.
From rt.restructuring.model.arrival Require Import arrival_curves.
From rt.restructuring.model.schedule Require Import
work_conserving priority_based.priorities priority_based.edf priority_based.preemption_aware.
From rt.restructuring.analysis.arrival Require Import workload_bound rbf.
From rt.restructuring.analysis.edf.rta Require Import nonpr_reg.response_time_bound.
(** Assume we have a fully non-preemptive model. *)
From rt.restructuring.model Require Import preemption.nonpreemptive.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq path fintype bigop.
(** * RTA for Fully Non-Preemptive FP Model *)
(** In this module we prove the RTA theorem for the fully non-preemptive EDF model. *)
Section RTAforFullyNonPreemptiveEDFModelwithArrivalCurves.
(** Consider any type of tasks ... *)
Context {Task : TaskType}.
Context `{TaskCost Task}.
Context `{TaskDeadline Task}.
(** ... and any type of jobs associated with these tasks. *)
Context {Job : JobType}.
Context `{JobTask Job Task}.
Context `{JobArrival Job}.
Context `{JobCost Job}.
(** For clarity, let's denote the relative deadline of a task as D. *)
Let D tsk := task_deadline tsk.
(** Consider the EDF policy that indicates a higher-or-equal priority relation. *)
Let EDF := EDF Task 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.
(** Consider an arbitrary task set ts, ... *)
Variable ts : list Task.
(** ... assume that all jobs come from this task set, ... *)
Hypothesis H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts.
(** ... and the cost of a job cannot be larger than the task cost. *)
Hypothesis H_job_cost_le_task_cost:
cost_of_jobs_from_arrival_sequence_le_task_cost arr_seq.
(** 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.
(** Let tsk be any task in ts that is to be analyzed. *)
Variable tsk : Task.
Hypothesis H_tsk_in_ts : tsk \in ts.
(** Next, consider any ideal non-preemptive uniprocessor schedule of this arrival sequence ... *)
Variable sched : schedule (ideal.processor_state Job).
Hypothesis H_nonpreemptive_sched : is_nonpreemptive_schedule sched.
Hypothesis H_jobs_come_from_arrival_sequence:
jobs_come_from_arrival_sequence sched arr_seq.
(** ... 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.
(** Assume we have sequential tasks, i.e, jobs from the
same task execute in the order of their arrival. *)
Hypothesis H_sequential_tasks : sequential_tasks sched.
(** Next, we assume that the schedule is a work-conserving schedule... *)
Hypothesis H_work_conserving : work_conserving arr_seq sched.
(** ... and the schedule respects the policy defined by the
job_preemptable function (i.e., jobs have bounded nonpreemptive
segments). *)
Hypothesis H_respects_policy : respects_policy_at_preemption_point arr_seq sched.
(** Let's define some local names for clarity. *)
Let response_time_bounded_by :=
task_response_time_bound arr_seq sched.
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.
(** We introduce the abbreviation "rbf" for the task request bound function,
which is defined as [task_cost(T) × max_arrivals(T,Δ)] for a task T. *)
Let rbf := task_request_bound_function.
(** Next, we introduce task_rbf as an abbreviation
for the task request bound function of task tsk. *)
Let task_rbf := rbf tsk.
(** Using the sum of individual request bound functions, we define the request bound
function of all tasks (total request bound function). *)
Let total_rbf := total_request_bound_function ts.
(** We also define a bound for the priority inversion caused by jobs with lower priority. *)
Let blocking_bound :=
\max_(tsk_o <- ts | (tsk_o != tsk) && (D tsk_o > D tsk))
(task_cost tsk_o - ε).
(** 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 + ε) + D tsk - D tsk_o) Δ).
(** 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_rbf L.
(** To reduce the time complexity of the analysis, recall the notion of search space. *)
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: nat.
Hypothesis H_R_is_maximum:
forall A,
is_in_search_space A ->
exists F,
A + F = blocking_bound + (task_rbf (A + ε) - (task_cost tsk - ε))
+ bound_on_total_hep_workload A (A + F) /\
F + (task_cost tsk - ε) <= R.
(** Now, we can leverage the results for the abstract model with bounded nonpreemptive segments
to establish a response-time bound for the more concrete model of fully nonpreemptive scheduling. *)
Theorem uniprocessor_response_time_bound_fully_nonpreemptive_edf:
response_time_bounded_by tsk R.
Proof.
case: (posnP (task_cost tsk)) => [ZERO|POS].
{ intros j ARR TSK.
have ZEROj: job_cost j = 0.
{ move: (H_job_cost_le_task_cost j ARR) => NEQ.
rewrite /job_cost_le_task_cost TSK ZERO in NEQ.
by apply/eqP; rewrite -leqn0.
}
by rewrite /job_response_time_bound /completed_by ZEROj.
}
eapply uniprocessor_response_time_bound_edf_with_bounded_nonpreemptive_segments with (L0 := L).
all: eauto 2 with basic_facts.
Qed.
End RTAforFullyNonPreemptiveEDFModelwithArrivalCurves.
From rt.util Require Import all.
From rt.restructuring.behavior Require Import all.
From rt.restructuring.analysis.basic_facts Require Import all.
From rt.restructuring.model Require Import job task workload processor.ideal readiness.basic.
From rt.restructuring.model.arrival Require Import arrival_curves.
From rt.restructuring.model.schedule Require Import
work_conserving priority_based.priorities priority_based.edf priority_based.preemption_aware.
From rt.restructuring.analysis.arrival Require Import workload_bound rbf.
From rt.restructuring.analysis.edf.rta Require Import nonpr_reg.response_time_bound.
(** Assume we have a fully preemptive model. *)
From rt.restructuring.model Require Import preemption.preemptive.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq path fintype bigop.
(** * RTA for Fully Preemptive EDF Model *)
(** In this section we prove the RTA theorem for the fully preemptive EDF model *)
Section RTAforFullyPreemptiveEDFModelwithArrivalCurves.
(** Consider any type of tasks ... *)
Context {Task : TaskType}.
Context `{TaskCost Task}.
Context `{TaskDeadline Task}.
(** ... and any type of jobs associated with these tasks. *)
Context {Job : JobType}.
Context `{JobTask Job Task}.
Context `{JobArrival Job}.
Context `{JobCost Job}.
(** For clarity, let's denote the relative deadline of a task as D. *)
Let D tsk := task_deadline tsk.
(** Consider the EDF policy that indicates a higher-or-equal priority relation. *)
Let EDF := EDF Task 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.
(** Consider an arbitrary task set ts, ... *)
Variable ts : list Task.
(** ... assume that all jobs come from this task set, ... *)
Hypothesis H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts.
(** ... and the cost of a job cannot be larger than the task cost. *)
Hypothesis H_job_cost_le_task_cost:
cost_of_jobs_from_arrival_sequence_le_task_cost arr_seq.
(** 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.
(** Let tsk be any task in ts that is to be analyzed. *)
Variable tsk : Task.
Hypothesis H_tsk_in_ts : tsk \in ts.
(** Next, consider any ideal uniprocessor schedule of the arrival sequence ... *)
Variable sched : schedule (ideal.processor_state Job).
Hypothesis H_jobs_come_from_arrival_sequence:
jobs_come_from_arrival_sequence sched arr_seq.
(** ... 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.
(** Assume we have sequential tasks, i.e, jobs from the
same task execute in the order of their arrival. *)
Hypothesis H_sequential_tasks : sequential_tasks sched.
(** Next, we assume that the schedule is a work-conserving schedule... *)
Hypothesis H_work_conserving : work_conserving arr_seq sched.
(** ... and the schedule respects the policy defined by the
job_preemptable function (i.e., jobs have bounded nonpreemptive
segments). *)
Hypothesis H_respects_policy : respects_policy_at_preemption_point arr_seq sched.
(** Let's define some local names for clarity. *)
Let response_time_bounded_by :=
task_response_time_bound arr_seq sched.
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.
(** We introduce the abbreviation "rbf" for the task request bound function,
which is defined as [task_cost(T) × max_arrivals(T,Δ)] for a task T. *)
Let rbf := task_request_bound_function.
(** Next, we introduce task_rbf as an abbreviation
for the task request bound function of task tsk. *)
Let task_rbf := rbf tsk.
(** Using the sum of individual request bound functions, we define the request bound
function of all tasks (total request bound function). *)
Let total_rbf := total_request_bound_function ts.
(** 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 + ε) + D tsk - D tsk_o) Δ).
(** 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_rbf L.
(** To reduce the time complexity of the analysis, recall the notion of search space. *)
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:
forall (A : duration),
is_in_search_space A ->
exists (F : duration),
A + F = task_rbf (A + ε) + bound_on_total_hep_workload A (A + F) /\
F <= R.
(** Now, we can leverage the results for the abstract model with bounded nonpreemptive segments
to establish a response-time bound for the more concrete model of fully preemptive scheduling. *)
Theorem uniprocessor_response_time_bound_fully_preemptive_edf:
response_time_bounded_by tsk R.
Proof.
have BLOCK: blocking_bound ts tsk = 0.
{ by rewrite /blocking_bound /parameters.task_max_nonpreemptive_segment
/preemptive.fully_preemptive_model subnn big1_eq. }
eapply uniprocessor_response_time_bound_edf_with_bounded_nonpreemptive_segments with (L0 := L) .
all: eauto 2 with basic_facts.
- move => A /andP [LT NEQ].
specialize (H_R_is_maximum A); feed H_R_is_maximum.
{ by apply/andP; split. }
move: H_R_is_maximum => [F [FIX BOUND]].
exists F; split.
+ by rewrite BLOCK add0n subnn subn0.
+ by rewrite subnn addn0.
Qed.
End RTAforFullyPreemptiveEDFModelwithArrivalCurves.
\ No newline at end of file
From rt.util Require Import all.
From rt.restructuring.behavior Require Import all.
From rt.restructuring.analysis.basic_facts Require Import all.
From rt.restructuring.model Require Import job task workload processor.ideal readiness.basic.
From rt.restructuring.model.arrival Require Import arrival_curves.
From rt.restructuring.model.preemption Require Import valid_model
job.parameters task.parameters rtc_threshold.valid_rtct.
From rt.restructuring.analysis.basic_facts.preemption Require Import
rtc_threshold.job_preemptable.
From rt.restructuring.analysis.facts Require Import priority_inversion_is_bounded.
From rt.restructuring.model.schedule Require Import
work_conserving priority_based.priorities priority_based.edf priority_based.preemption_aware.
From rt.restructuring.analysis.arrival Require Import workload_bound rbf.
From rt.restructuring.analysis.edf Require Export rta.response_time_bound.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq path fintype bigop.
(** * RTA for EDF-schedulers with Bounded Non-Preemprive Segments *)
(** In this section we instantiate the Abstract RTA for EDF-schedulers
with Bounded Priority Inversion to EDF-schedulers for ideal
uni-processor model of real-time tasks with arbitrary
arrival models _and_ bounded non-preemprive segments. *)
(** Recall that Abstract RTA for EDF-schedulers with Bounded Priority
Inversion does not specify the cause of priority inversion. In
this section, we prove that the priority inversion caused by
execution of non-preemptive segments is bounded. Thus the Abstract
RTA for EDF-schedulers is applicable to this instantiation. *)
Section RTAforEDFwithBoundedNonpreemptiveSegmentsWithArrivalCurves.
(** Consider any type of tasks ... *)
Context {Task : TaskType}.
Context `{TaskCost Task}.
Context `{TaskDeadline Task}.
Context `{TaskRunToCompletionThreshold Task}.
Context `{TaskMaxNonpreemptiveSegment Task}.
(** ... and any type of jobs associated with these tasks. *)
Context {Job : JobType}.
Context `{JobTask Job Task}.
Context `{JobArrival Job}.
Context `{JobCost Job}.
(** For clarity, let's denote the relative deadline of a task as D. *)
Let D tsk := task_deadline tsk.
(** Consider the EDF policy that indicates a higher-or-equal priority relation.
Note that we do not relate the EDF policy with the scheduler. However, we
define functions for Interference and Interfering Workload that actively use
the concept of priorities. *)
Let EDF := EDF Task 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.
(** Next, consider any ideal uniprocessor schedule 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.
(** ... 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.
(** In addition, we assume the existence of a function maping jobs
to theirs preemption points ... *)
Context `{JobPreemptable Job}.
(** ... and assume that it defines a valid preemption
model with bounded nonpreemptive segments. *)
Hypothesis H_valid_model_with_bounded_nonpreemptive_segments:
valid_model_with_bounded_nonpreemptive_segments
arr_seq sched.
(** Assume we have sequential tasks, i.e, jobs from the
same task execute in the order of their arrival. *)
Hypothesis H_sequential_tasks : sequential_tasks sched.
(** Next, we assume that the schedule is a work-conserving schedule... *)
Hypothesis H_work_conserving : work_conserving arr_seq sched.
(** ... and the schedule respects the policy defined by thejob_preemptable
function (i.e., jobs have bounded nonpreemptive segments). *)
Hypothesis H_respects_policy : respects_policy_at_preemption_point arr_seq sched.
(** Consider an arbitrary task set ts, ... *)
Variable ts : list Task.
(** ... assume that all jobs come from the task set, ... *)
Hypothesis H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts.
(** ... and the cost of a job cannot be larger than the task cost. *)
Hypothesis H_job_cost_le_task_cost:
cost_of_jobs_from_arrival_sequence_le_task_cost arr_seq.
(** 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.
(** Let tsk be any task in ts that is to be analyzed. *)
Variable tsk : Task.
Hypothesis H_tsk_in_ts : tsk \in ts.
(** Consider a valid preemption model... *)
Hypothesis H_valid_preemption_model:
valid_preemption_model arr_seq sched.
(** ...and a valid task run-to-completion threshold function. That is,
[task_run_to_completion_threshold tsk] is (1) no bigger than tsk's
cost, (2) for any job of task tsk job_run_to_completion_threshold
is bounded by task_run_to_completion_threshold. *)
Hypothesis H_valid_run_to_completion_threshold:
valid_task_run_to_completion_threshold arr_seq tsk.
(** We introduce as an abbreviation "rbf" for the task request bound function,
which is defined as [task_cost(T) × max_arrivals(T,Δ)] for a task T. *)
Let rbf := task_request_bound_function.
(** Next, we introduce task_rbf as an abbreviation for the task
request bound function of task tsk. *)
Let task_rbf := rbf tsk.
(** Using the sum of individual request bound functions, we define the request bound
function of all tasks (total request bound function). *)
Let total_rbf := total_request_bound_function ts.
(** 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 + ε) + D tsk - D tsk_o) Δ).
(** Let's define some local names for clarity. *)
Let max_length_of_priority_inversion :=
max_length_of_priority_inversion arr_seq EDF.
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 response_time_bounded_by := task_response_time_bound arr_seq sched.
(** We also define a bound for the priority inversion caused by jobs with lower priority. *)
Definition blocking_bound :=
\max_(tsk_o <- ts | (tsk_o != tsk) && (D tsk < D tsk_o))
(task_max_nonpreemptive_segment tsk_o - ε).
(** ** Priority inversion is bounded *)
(** In this section, we prove that a priority inversion for task tsk is bounded by
the maximum length of nonpreemtive segments among the tasks with lower priority. *)
Section PriorityInversionIsBounded.
(** First, we prove that the maximum length of a priority
inversion of job j is bounded by the maximum length of a
nonpreemptive section of a task with lower-priority task
(i.e., the blocking term). *)
Lemma priority_inversion_is_bounded_by_blocking:
forall j t,
arrives_in arr_seq j ->
job_task j = tsk ->
t <= job_arrival j ->
max_length_of_priority_inversion j t <= blocking_bound.
Proof.
intros j t ARR TSK LE; unfold max_length_of_priority_inversion, blocking_bound.
apply leq_trans with
(\max_(j_lp <- arrivals_between arr_seq 0 t | ~~ EDF j_lp j)
(task_max_nonpreemptive_segment (job_task j_lp) - ε)).
- apply leq_big_max.
intros j' JINB NOTHEP.
rewrite leq_sub2r //.
apply in_arrivals_implies_arrived in JINB.
by apply H_valid_model_with_bounded_nonpreemptive_segments.
- apply /bigmax_leq_seqP.
intros j' JINB NOTHEP.
apply leq_bigmax_cond_seq with (i0 := (job_task j')) (F := fun tsk => task_max_nonpreemptive_segment tsk - 1).
{ apply H_all_jobs_from_taskset.
apply mem_bigcat_nat_exists in JINB.
by inversion JINB as [ta' [JIN' _]]; exists ta'. }
{ have NINTSK: job_task j' != tsk.
{ apply/eqP; intros TSKj'.
rewrite /EDF -ltnNge in NOTHEP.
rewrite /job_deadline /job_deadline.job_deadline_from_task_deadline in NOTHEP.
rewrite TSKj' TSK ltn_add2r in NOTHEP.
move: NOTHEP; rewrite ltnNge; move => /negP T; apply: T.
apply leq_trans with t; last by done.
eapply in_arrivals_implies_arrived_between in JINB; last by eauto 2.
move: JINB; move => /andP [_ T].
by apply ltnW.
}
apply/andP; split; first by done.
rewrite /EDF -ltnNge in NOTHEP.
rewrite -TSK.
have ARRLE: job_arrival j' < job_arrival j.
{ apply leq_trans with t; last by done.
eapply in_arrivals_implies_arrived_between in JINB; last by eauto 2.
by move: JINB; move => /andP [_ T].
}
rewrite /job_deadline /job_deadline.job_deadline_from_task_deadline in NOTHEP.
rewrite /D; ssromega.
}
Qed.
(** Using the lemma above, we prove that the priority inversion of the task is bounded by
the maximum length of a nonpreemptive section of lower-priority tasks. *)
Lemma priority_inversion_is_bounded:
priority_inversion_is_bounded_by arr_seq sched _ tsk blocking_bound.
Proof.
move => j ARR TSK POS t1 t2 PREF; move: (PREF) => [_ [_ [_ /andP [T _]]]].
destruct (leqP (t2 - t1) blocking_bound) as [NEQ|NEQ].
{ apply leq_trans with (t2 - t1); last by done.
rewrite /cumulative_priority_inversion /is_priority_inversion.
rewrite -[X in _ <= X]addn0 -[t2 - t1]mul1n -iter_addn -big_const_nat.
rewrite leq_sum //.
intros t _; case: (sched t); last by done.
by intros s; case: edf.EDF.
}
edestruct @preemption_time_exists as [ppt [PPT NEQ2]]; eauto 2 with basic_facts.
move: NEQ2 => /andP [GE LE].
apply leq_trans with (cumulative_priority_inversion sched _ j t1 ppt);
last apply leq_trans with (ppt - t1).
- rewrite /cumulative_priority_inversion /is_priority_inversion.
rewrite (@big_cat_nat _ _ _ ppt) //=; last first.
{ rewrite ltn_subRL in NEQ.
apply leq_trans with (t1 + blocking_bound); last by apply ltnW.
apply leq_trans with (t1 + max_length_of_priority_inversion j t1); first by done.
by rewrite leq_add2l; eapply priority_inversion_is_bounded_by_blocking; eauto 2. }
rewrite -[X in _ <= X]addn0 leq_add2l leqn0.
rewrite big_nat_cond big1 //; move => t /andP [/andP [GEt LTt] _ ].
case SCHED: (sched t) => [s | ]; last by done.
edestruct @not_quiet_implies_exists_scheduled_hp_job
with (K := ppt - t1) (t := t) as [j_hp [ARRB [HP SCHEDHP]]]; eauto 2 with basic_facts.
{ exists ppt; split. by done. by rewrite subnKC //; apply/andP; split. }
{ by rewrite subnKC //; apply/andP; split. }
apply/eqP; rewrite eqb0 Bool.negb_involutive.
enough (EQ : s = j_hp); first by subst.
move: SCHED => /eqP SCHED; rewrite -scheduled_at_def in SCHED.
by eapply ideal_proc_model_is_a_uniprocessor_model; [exact SCHED | exact SCHEDHP].
- rewrite /cumulative_priority_inversion /is_priority_inversion.
rewrite -[X in _ <= X]addn0 -[ppt - t1]mul1n -iter_addn -big_const_nat.
rewrite leq_sum //.
intros t _; case: (sched t); last by done.
by intros s; case: edf.EDF.
- rewrite leq_subLR.
apply leq_trans with (t1 + max_length_of_priority_inversion j t1); first by done.
rewrite leq_add2l; eapply priority_inversion_is_bounded_by_blocking; eauto 2.
Qed.
End PriorityInversionIsBounded.
(** ** Response-Time Bound *)
(** In this section, we prove that the maximum among the solutions of the response-time
bound recurrence is a response-time bound for tsk. *)
Section ResponseTimeBound.
(** 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_rbf L.
(** To reduce the time complexity of the analysis, recall the notion of search space. *)
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:
forall (A : duration),
is_in_search_space A ->
exists (F : duration),
A + F = blocking_bound
+ (task_rbf (A + ε) - (task_cost tsk - task_run_to_completion_threshold tsk))
+ bound_on_total_hep_workload A (A + F) /\
F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R.
(** Then, using the results for the general RTA for EDF-schedulers, we establish a
response-time bound for the more concrete model of bounded nonpreemptive segments.
Note that in case of the general RTA for EDF-schedulers, we just _assume_ that
the priority inversion is bounded. In this module we provide the preemption model
with bounded nonpreemptive segments and _prove_ that the priority inversion is
bounded. *)
Theorem uniprocessor_response_time_bound_edf_with_bounded_nonpreemptive_segments:
response_time_bounded_by tsk R.
Proof.
eapply uniprocessor_response_time_bound_edf; eauto 2.
by apply priority_inversion_is_bounded.
Qed.
End ResponseTimeBound.
End RTAforEDFwithBoundedNonpreemptiveSegmentsWithArrivalCurves.
This diff is collapsed.
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment