Library prosa.results.rta.arm.edf.limited_preemptive
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.edf.
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.average.
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.edf.
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.average.
RTA for EDF Scheduling with Fixed Preemption Points on Uniprocessors under the Average 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} `{TaskDeadline Task}
`{MaxArrivals Task} `{TaskPreemptionPoints Task}.
`{MaxArrivals Task} `{TaskPreemptionPoints 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 job's preemption points job_preemptive_points. 
  Context {Job : JobType} `{JobTask Job Task} `{JobCost Job} `{JobArrival Job}
`{JobPreemptionPoints Job}.
`{JobPreemptionPoints Job}.
We assume that jobs are limited-preemptive. 
  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 a model with fixed preemption points. I.e., each task is divided
      into a number of non-preemptive segments by inserting statically
      predefined preemption points. 
  Hypothesis H_valid_model_with_fixed_preemption_points :
valid_fixed_preemption_points_model arr_seq ts.
valid_fixed_preemption_points_model arr_seq ts.
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_schedule_with_limited_preemptions :
schedule_respects_preemption_model arr_seq sched.
Hypothesis H_valid_schedule : valid_schedule sched arr_seq.
Hypothesis H_work_conserving : work_conserving arr_seq sched.
Hypothesis H_schedule_with_limited_preemptions :
schedule_respects_preemption_model arr_seq sched.
We assume that the schedule respects the given EDF scheduling policy. 
Average Resource Model
[t1, t2), and given a
      resource period Π, a resource allocation time Θ, and a supply delay
      ν, the processor provides at least (t2 - t1 - ν) ⋅ Θ / Π units of
      supply. Intuitively, this means that on _average, the processor
      delivers Θ units of output every Π units of time, while the
      distribution of supply is not ideal and is subject to fluctuations bounded
      by ν. Furthermore, let arm_sbf Π Θ ν denote the SBF, which, as proven
      in prosa.analysis.facts.model.sbf.average, is a valid SBF. 
  Variables Π Θ ν : duration.
Hypothesis H_average_resource_model : average_resource_model Π Θ ν sched.
Hypothesis H_average_resource_model : average_resource_model Π Θ ν sched.
Maximum Length of a Busy Interval
  Definition busy_window_recurrence_solution (L : duration) :=
L > 0
∧ arm_sbf Π Θ ν L ≥ total_request_bound_function ts L
∧ arm_sbf Π Θ ν L ≥ longest_busy_interval_with_pi ts tsk.
L > 0
∧ arm_sbf Π Θ ν L ≥ total_request_bound_function ts L
∧ arm_sbf Π Θ ν L ≥ longest_busy_interval_with_pi ts tsk.
Response-Time Bound
  Definition rta_recurrence_solution L R :=
∀ (A : duration),
is_in_search_space ts tsk L A →
∃ (F : duration),
arm_sbf Π Θ ν F ≥ blocking_bound ts tsk A
+ (task_request_bound_function tsk (A + ε) - (task_last_nonpr_segment tsk - ε))
+ bound_on_athep_workload ts tsk A F
∧ arm_sbf Π Θ ν (A + R) ≥ arm_sbf Π Θ ν F + (task_last_nonpr_segment tsk - ε)
∧ A + R ≥ F.
∀ (A : duration),
is_in_search_space ts tsk L A →
∃ (F : duration),
arm_sbf Π Θ ν F ≥ blocking_bound ts tsk A
+ (task_request_bound_function tsk (A + ε) - (task_last_nonpr_segment tsk - ε))
+ bound_on_athep_workload ts tsk A F
∧ arm_sbf Π Θ ν (A + R) ≥ arm_sbf Π Θ ν F + (task_last_nonpr_segment tsk - ε)
∧ 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 EDF scheduling with limited preemptions on a unit-speed
      uniprocessor under the average resource model. 
  Theorem uniprocessor_response_time_bound_limited_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_SOL1 BW_SOL2]] R SOL js ARRs TSKs.
have [ZERO|POS] := posnP (job_cost js); first by rewrite /job_response_time_bound /completed_by ZERO.
have VAL1 : valid_preemption_model arr_seq sched
by apply valid_fixed_preemption_points_model_lemma, H_valid_model_with_fixed_preemption_points.
have VBSBF : valid_busy_sbf arr_seq sched tsk (arm_sbf Π Θ ν).
{ by apply: valid_pred_sbf_switch_predicate; last apply arm_sbf_valid. }
have US : unit_supply_bound_function (arm_sbf Π Θ ν) by exact: arm_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_seq with (L := L) (SBF := arm_sbf Π Θ ν) ⇒ //.
- exact: instantiated_i_and_w_are_coherent_with_schedule.
- exact: EDF_implies_sequential_tasks.
- exact: instantiated_interference_and_workload_consistent_with_sequential_tasks.
- apply: busy_intervals_are_bounded_rs_edf ⇒ //.
by apply: instantiated_i_and_w_are_coherent_with_schedule.
- apply: valid_pred_sbf_switch_predicate; last first.
+ by apply arm_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 [FIX1 [FIX2 FIX3]]; ∃ F; split ⇒ //; split.
× erewrite last_segment_eq_cost_minus_rtct; [ | eauto | eauto ].
by rewrite /task_intra_IBF; lia.
× by erewrite last_segment_eq_cost_minus_rtct; [ | eauto | eauto ]; lia.
Qed.
End RTAforLimitedPreemptiveEDFModelwithArrivalCurves.
∀ (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_SOL1 BW_SOL2]] R SOL js ARRs TSKs.
have [ZERO|POS] := posnP (job_cost js); first by rewrite /job_response_time_bound /completed_by ZERO.
have VAL1 : valid_preemption_model arr_seq sched
by apply valid_fixed_preemption_points_model_lemma, H_valid_model_with_fixed_preemption_points.
have VBSBF : valid_busy_sbf arr_seq sched tsk (arm_sbf Π Θ ν).
{ by apply: valid_pred_sbf_switch_predicate; last apply arm_sbf_valid. }
have US : unit_supply_bound_function (arm_sbf Π Θ ν) by exact: arm_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_seq with (L := L) (SBF := arm_sbf Π Θ ν) ⇒ //.
- exact: instantiated_i_and_w_are_coherent_with_schedule.
- exact: EDF_implies_sequential_tasks.
- exact: instantiated_interference_and_workload_consistent_with_sequential_tasks.
- apply: busy_intervals_are_bounded_rs_edf ⇒ //.
by apply: instantiated_i_and_w_are_coherent_with_schedule.
- apply: valid_pred_sbf_switch_predicate; last first.
+ by apply arm_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 [FIX1 [FIX2 FIX3]]; ∃ F; split ⇒ //; split.
× erewrite last_segment_eq_cost_minus_rtct; [ | eauto | eauto ].
by rewrite /task_intra_IBF; lia.
× by erewrite last_segment_eq_cost_minus_rtct; [ | eauto | eauto ]; lia.
Qed.
End RTAforLimitedPreemptiveEDFModelwithArrivalCurves.