Library prosa.classic.implementation.uni.jitter.fp_rta_example
Require Import prosa.classic.util.all.
Require Import prosa.classic.model.priority.
Require Import prosa.classic.model.arrival.basic.task.
Require Import prosa.classic.model.schedule.uni.schedulability.
Require Import prosa.classic.model.arrival.jitter.job.
Require Import prosa.classic.model.schedule.uni.jitter.schedule.
Require Import prosa.classic.analysis.uni.jitter.workload_bound_fp
prosa.classic.analysis.uni.jitter.fp_rta_comp.
Require Import prosa.classic.implementation.uni.jitter.job
prosa.classic.implementation.uni.jitter.task
prosa.classic.implementation.uni.jitter.arrival_sequence
prosa.classic.implementation.uni.jitter.schedule.
From mathcomp Require Import ssreflect ssrbool ssrnat eqtype seq bigop div.
Module ResponseTimeAnalysisFP.
Import JobWithJitter UniprocessorScheduleWithJitter SporadicTaskset Priority Schedulability
WorkloadBoundFP ResponseTimeIterationFP.
Import ConcreteJob ConcreteTask ConcreteArrivalSequence ConcreteScheduler.
(* In this section, we test the jitter-aware FP RTA on a simple task set
to show that the theorems contain no contradictory assumptions. *)
Section ExampleRTA.
Let tsk1 := {| task_id := 1; task_cost := 1; task_period := 5;
task_deadline := 6; task_jitter := 1 |}.
Let tsk2 := {| task_id := 2; task_cost := 1; task_period := 5;
task_deadline := 6; task_jitter := 0|}.
Let tsk3 := {| task_id := 3; task_cost := 1; task_period := 6;
task_deadline := 6; task_jitter := 1|}.
(* Let ts be a task set containing these three tasks, ... *)
Program Let ts := Build_set [:: tsk1; tsk2; tsk3] _.
(* ...which can be shown to have valid parameters. *)
Remark ts_has_positive_costs:
∀ tsk, tsk \in ts → task_cost tsk > 0.
Remark ts_has_positive_periods:
∀ tsk, tsk \in ts → task_period tsk > 0.
(* Next, recall the FP RTA schedulability test using RM as the FP policy. *)
Let RTA_claimed_bounds :=
fp_claimed_bounds task_cost task_period task_deadline task_jitter (RM task_period).
Let schedulability_test :=
fp_schedulable task_cost task_period task_deadline task_jitter (RM task_period).
(* First, we show that the schedulability test returns the following bounds.
(Note that although we check the task jitter, these values only include the
response-time bounds.) *)
Fact RTA_yields_these_bounds :
RTA_claimed_bounds ts = Some [:: (tsk1, 2); (tsk2, 2); (tsk3, 3)].
(* Since the response-time bounds plus task jitter are no larger than task deadlines,
the schedulability test indeed returns true. *)
Fact schedulability_test_succeeds :
schedulability_test ts = true.
(* Now, let's show that the task set is schedulable. *)
(* Let arr_seq be the periodic arrival sequence from ts... *)
Let arr_seq := periodic_arrival_sequence ts.
(* ...subject to rate-monotonic priority. *)
Let higher_eq_priority := FP_to_JLDP job_task (RM task_period).
(* Assume that jobs have jitter that is no larger than the jitter of their tasks. *)
Variable job_jitter: concrete_job → time.
Hypothesis H_jitter_is_bounded:
∀ j,
arrives_in arr_seq j →
job_jitter_leq_task_jitter task_jitter job_jitter job_task j.
(* Next, let sched be the jitter-aware RM schedule with those jitter values. *)
Let sched := scheduler job_arrival job_cost job_jitter arr_seq higher_eq_priority.
(* To conclude, based on the definition of deadline miss,... *)
Let no_deadline_missed_by :=
task_misses_no_deadline job_arrival job_cost job_deadline job_task arr_seq sched.
(* ...we use the result of the jitter-aware FP RTA to conclude that
no task misses its deadline. *)
Corollary ts_is_schedulable:
∀ tsk,
tsk \in ts →
no_deadline_missed_by tsk.
End ExampleRTA.
End ResponseTimeAnalysisFP.
Require Import prosa.classic.model.priority.
Require Import prosa.classic.model.arrival.basic.task.
Require Import prosa.classic.model.schedule.uni.schedulability.
Require Import prosa.classic.model.arrival.jitter.job.
Require Import prosa.classic.model.schedule.uni.jitter.schedule.
Require Import prosa.classic.analysis.uni.jitter.workload_bound_fp
prosa.classic.analysis.uni.jitter.fp_rta_comp.
Require Import prosa.classic.implementation.uni.jitter.job
prosa.classic.implementation.uni.jitter.task
prosa.classic.implementation.uni.jitter.arrival_sequence
prosa.classic.implementation.uni.jitter.schedule.
From mathcomp Require Import ssreflect ssrbool ssrnat eqtype seq bigop div.
Module ResponseTimeAnalysisFP.
Import JobWithJitter UniprocessorScheduleWithJitter SporadicTaskset Priority Schedulability
WorkloadBoundFP ResponseTimeIterationFP.
Import ConcreteJob ConcreteTask ConcreteArrivalSequence ConcreteScheduler.
(* In this section, we test the jitter-aware FP RTA on a simple task set
to show that the theorems contain no contradictory assumptions. *)
Section ExampleRTA.
Let tsk1 := {| task_id := 1; task_cost := 1; task_period := 5;
task_deadline := 6; task_jitter := 1 |}.
Let tsk2 := {| task_id := 2; task_cost := 1; task_period := 5;
task_deadline := 6; task_jitter := 0|}.
Let tsk3 := {| task_id := 3; task_cost := 1; task_period := 6;
task_deadline := 6; task_jitter := 1|}.
(* Let ts be a task set containing these three tasks, ... *)
Program Let ts := Build_set [:: tsk1; tsk2; tsk3] _.
(* ...which can be shown to have valid parameters. *)
Remark ts_has_positive_costs:
∀ tsk, tsk \in ts → task_cost tsk > 0.
Remark ts_has_positive_periods:
∀ tsk, tsk \in ts → task_period tsk > 0.
(* Next, recall the FP RTA schedulability test using RM as the FP policy. *)
Let RTA_claimed_bounds :=
fp_claimed_bounds task_cost task_period task_deadline task_jitter (RM task_period).
Let schedulability_test :=
fp_schedulable task_cost task_period task_deadline task_jitter (RM task_period).
(* First, we show that the schedulability test returns the following bounds.
(Note that although we check the task jitter, these values only include the
response-time bounds.) *)
Fact RTA_yields_these_bounds :
RTA_claimed_bounds ts = Some [:: (tsk1, 2); (tsk2, 2); (tsk3, 3)].
(* Since the response-time bounds plus task jitter are no larger than task deadlines,
the schedulability test indeed returns true. *)
Fact schedulability_test_succeeds :
schedulability_test ts = true.
(* Now, let's show that the task set is schedulable. *)
(* Let arr_seq be the periodic arrival sequence from ts... *)
Let arr_seq := periodic_arrival_sequence ts.
(* ...subject to rate-monotonic priority. *)
Let higher_eq_priority := FP_to_JLDP job_task (RM task_period).
(* Assume that jobs have jitter that is no larger than the jitter of their tasks. *)
Variable job_jitter: concrete_job → time.
Hypothesis H_jitter_is_bounded:
∀ j,
arrives_in arr_seq j →
job_jitter_leq_task_jitter task_jitter job_jitter job_task j.
(* Next, let sched be the jitter-aware RM schedule with those jitter values. *)
Let sched := scheduler job_arrival job_cost job_jitter arr_seq higher_eq_priority.
(* To conclude, based on the definition of deadline miss,... *)
Let no_deadline_missed_by :=
task_misses_no_deadline job_arrival job_cost job_deadline job_task arr_seq sched.
(* ...we use the result of the jitter-aware FP RTA to conclude that
no task misses its deadline. *)
Corollary ts_is_schedulable:
∀ tsk,
tsk \in ts →
no_deadline_missed_by tsk.
End ExampleRTA.
End ResponseTimeAnalysisFP.