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

Add platform for models with bounded nonpreemptive segments

parent 13dc6eaf
No related branches found
No related tags found
1 merge request!9Complete proof of Abstract RTA
Require Import rt.util.all.
Require Import rt.model.arrival.basic.job
rt.model.arrival.basic.task
rt.model.priority
rt.model.arrival.basic.task_arrival.
Require Import rt.model.schedule.uni.schedule
rt.model.schedule.uni.service
rt.model.schedule.uni.basic.platform.
Require Import rt.model.schedule.uni.nonpreemptive.schedule.
From mathcomp Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq fintype bigop.
(** * Platform with limited preemptions *)
(** In this module we introduce the notion of whether a job can be preempted at a given time
(using a predicate can_be_preempted). In addition, we provide instantiations of the
predicate for various preemption models. *)
Module LimitedPreemptionPlatform.
Import Epsilon Job SporadicTaskset UniprocessorSchedule Priority Service.
(* In this section, we define a processor platform with limited preemptions. *)
Section Properties.
Context {Task: eqType}.
Variable task_cost: Task -> time.
Context {Job: eqType}.
Variable job_arrival: Job -> time.
Variable job_cost: Job -> time.
Variable job_task: Job -> Task.
(* Consider any job arrival sequence... *)
Variable arr_seq: arrival_sequence Job.
(* ...and any uniprocessor schedule of these jobs. *)
Variable sched: schedule Job.
(* For simplicity, let's define some local names. *)
Let job_pending := pending job_arrival job_cost sched.
Let job_completed_by := completed_by job_cost sched.
Let job_scheduled_at := scheduled_at sched.
(* First, we define the notion of a preemption time. *)
Section PreemptionTime.
(* Let can_be_preempted be a function that maps a job j and the progress of j
at some time instant t to a boolean value, i.e., true if job j can be
preempted at this point of execution and false otherwise. *)
Variable can_be_preempted: Job -> time -> bool.
(* We say that a time instant t is a preemption time iff there's no job currently
scheduled at t that cannot be preempted (according to the predicate). *)
Definition preemption_time (t: time) :=
if sched t is Some j then
can_be_preempted j (service sched j t)
else true.
(* Since the notion of preemption time is based on an user-provided
predicate (variable can_be_preempted), it does not guarantee that
the scheduler will enforce correct behavior. For that, we must
define additional predicates. *)
Section CorrectPreemptionModel.
(* First, if a job j is not preemptive at some time instant t,
then j must be scheduled at time t. *)
Definition not_preemptive_implies_scheduled (j: Job) :=
forall t,
~~ can_be_preempted j (service sched j t) ->
job_scheduled_at j t.
(* A job can start its execution only from a preemption point. *)
Definition execution_starts_with_preemption_point (j: Job) :=
forall prt,
~~ job_scheduled_at j prt ->
job_scheduled_at j prt.+1 ->
can_be_preempted j (service sched j prt.+1).
(* We say that a model is a correct preemption model if both
definitions given above are satisfied for any job. *)
Definition correct_preemption_model :=
forall j,
arrives_in arr_seq j ->
not_preemptive_implies_scheduled j
/\ execution_starts_with_preemption_point j.
End CorrectPreemptionModel.
(* Note that for analysis purposes, it is important that the distance
between preemption points of a job is bounded. To ensure that, we
define next the model of bounded nonpreemptive segment. *)
Section ModelWithBoundedNonpreemptiveRegions.
(* We require that a job has to be executed at least one time instant
in order to reach a nonpreemptive segment. *)
Definition job_cannot_become_nonpreemptive_before_execution (j: Job) :=
can_be_preempted j 0.
(* And vice versa, a job cannot remain nonpreemptive after its completion. *)
Definition job_cannot_be_nonpreemptive_after_completion (j: Job) :=
can_be_preempted j (job_cost j).
(* Consider a function that maps a job to the length of
its maximal nonpreemptive segment. *)
Variable job_max_nps: Job -> time.
(* And a function task_max_nps... *)
Variable task_max_nps: Task -> time.
(* ...that gives an upper bound for values of the function job_max_nps. *)
Definition job_max_nonpreemptive_segment_le_task_max_nonpreemptive_segment (j: Job) :=
arrives_in arr_seq j ->
job_max_nps j <= task_max_nps (job_task j).
(* Next, we say that all the segments of a job j have bounded length iff for any
progress progr of job j there exists a preemption point preeemption_point such that
[progr <= preemption_point <= progr + (job_max_nps j - ε)]. That is, in any time
interval of length [job_max_nps j], there exists a preeemption point which
lies in this interval. *)
Definition nonpreemptive_regions_have_bounded_length (j: Job) :=
forall progr,
0 <= progr <= job_cost j ->
exists preemption_point,
progr <= preemption_point <= progr + (job_max_nps j - ε) /\
can_be_preempted j preemption_point.
(* Finally, we say that the schedule enforces bounded nonpreemptive segments
iff the predicate can_be_preempted satisfies the two conditions above. *)
Definition model_with_bounded_nonpreemptive_segments :=
forall j,
arrives_in arr_seq j ->
job_cannot_become_nonpreemptive_before_execution j
/\ job_cannot_be_nonpreemptive_after_completion j
/\ job_max_nonpreemptive_segment_le_task_max_nonpreemptive_segment j
/\ nonpreemptive_regions_have_bounded_length j.
End ModelWithBoundedNonpreemptiveRegions.
(* In this section we prove a few basic properties of the can_be_preempted predicate. *)
Section Lemmas.
Variable job_max_nps: Job -> time.
Variable task_max_nps: Task -> time.
(* Consider the correct model with bounded nonpreemptive segments. *)
Hypothesis H_correct_preemption_model: correct_preemption_model.
Hypothesis H_model_with_bounded_np_segments:
model_with_bounded_nonpreemptive_segments job_max_nps task_max_nps.
(* Assume jobs come from some arrival sequence. *)
Hypothesis H_jobs_come_from_arrival_sequence:
jobs_come_from_arrival_sequence sched arr_seq.
(* Then, we can show that time 0 is a preemption time. *)
Lemma zero_is_pt: preemption_time 0.
Proof.
unfold preemption_time.
case SCHED: (sched 0) => [j | ]; last by done.
move: (SCHED) => /eqP ARR.
apply H_jobs_come_from_arrival_sequence in ARR.
rewrite /service /service_during big_geq; last by done.
by move: (H_model_with_bounded_np_segments j ARR) => [PP _]; apply PP.
Qed.
(* Also, we show that the first instant of execution is a preemption time. *)
Lemma first_moment_is_pt:
forall j prt,
arrives_in arr_seq j ->
~~ job_scheduled_at j prt ->
job_scheduled_at j prt.+1 ->
preemption_time prt.+1.
Proof.
intros s pt ARR NSCHED SCHED.
unfold preemption_time.
move: (SCHED) => /eqP SCHED2; rewrite SCHED2; clear SCHED2.
by move: (H_correct_preemption_model s ARR) => [_ FHF]; auto.
Qed.
End Lemmas.
End PreemptionTime.
(* Next, we define properties related to execution. *)
Section Execution.
(* Similarly to preemptive scheduling, we say that the schedule is
work-conserving iff whenever a job is backlogged, the processor
is always busy scheduling another job. *)
(* Imported from the preemptive schedule. *)
Definition work_conserving := Platform.work_conserving job_cost.
End Execution.
(* Next, we define properties related to FP scheduling. *)
Section FP.
(* Consider any preemption model. *)
Variable preemption_model: Job -> time -> bool.
(* We say that an FP policy...*)
Variable higher_eq_priority: FP_policy Task.
(* ...is respected by the schedule iff, at every preemption point,
a scheduled task has higher (or same) priority than (as)
any backlogged task. *)
Definition respects_FP_policy_at_preemption_point :=
forall j j_hp t,
preemption_time preemption_model t ->
arrives_in arr_seq j ->
backlogged job_arrival job_cost sched j t ->
scheduled_at sched j_hp t ->
higher_eq_priority (job_task j_hp) (job_task j).
End FP.
(* Next, we define properties related to JLFP policies. *)
Section JLFP.
(* Consider a scheduling model. *)
Variable preemption_model: Job -> time -> bool.
(* We say that a JLFP policy ...*)
Variable higher_eq_priority: JLFP_policy Job.
(* ...is respected by the schedule iff, at every preemption point,
a scheduled task has higher (or same) priority than (as)
any backlogged task. *)
Definition respects_JLFP_policy_at_preemption_point :=
forall j j_hp t,
preemption_time preemption_model t ->
arrives_in arr_seq j ->
backlogged job_arrival job_cost sched j t ->
scheduled_at sched j_hp t ->
higher_eq_priority j_hp j.
End JLFP.
End Properties.
End LimitedPreemptionPlatform.
\ No newline at end of file
This diff is collapsed.
Require Import rt.util.all.
Require Import rt.model.arrival.basic.job
rt.model.arrival.basic.task
rt.model.priority
rt.model.arrival.basic.task_arrival.
Require Import rt.model.schedule.uni.schedule
rt.model.schedule.uni.service
rt.model.schedule.uni.basic.platform.
Require Import rt.model.schedule.uni.nonpreemptive.schedule.
Require Export rt.model.schedule.uni.limited.platform.definitions.
From mathcomp Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq fintype bigop.
(** * Platform for fully nonpreemptive model *)
(** In module uni.limited.platform we introduce the notion of whether a job can be preempted
at a given time (using a predicate can_be_preempted). In this section, we instantiate
can_be_preempted for the fully nonpreemptive model and prove its correctness. *)
Module FullyNonPreemptivePlatform.
Import Epsilon Job SporadicTaskset UniprocessorSchedule Priority
Service LimitedPreemptionPlatform.
Section FullyNonPreemptiveModel.
Context {Task: eqType}.
Variable task_cost: Task -> time.
Context {Job: eqType}.
Variable job_arrival: Job -> time.
Variable job_cost: Job -> time.
Variable job_task: Job -> Task.
(* Consider any arrival sequence with consistent, non-duplicate arrivals. *)
Variable arr_seq: arrival_sequence Job.
Hypothesis H_arrival_times_are_consistent: arrival_times_are_consistent job_arrival arr_seq.
Hypothesis H_arr_seq_is_a_set: arrival_sequence_is_a_set arr_seq.
(* Next, consider any uniprocessor nonpreemptive schedule of this arrival sequence...*)
Variable sched: schedule Job.
Hypothesis H_jobs_come_from_arrival_sequence: jobs_come_from_arrival_sequence sched arr_seq.
Hypothesis H_nonpreemptive_sched:
NonpreemptiveSchedule.is_nonpreemptive_schedule job_cost sched.
(* ... where jobs do not execute before their arrival nor after completion. *)
Hypothesis H_jobs_must_arrive_to_execute: jobs_must_arrive_to_execute job_arrival sched.
Hypothesis H_completed_jobs_dont_execute: completed_jobs_dont_execute job_cost sched.
(* For simplicity, let's define some local names. *)
Let job_pending := pending job_arrival job_cost sched.
Let job_completed_by := completed_by job_cost sched.
Let job_scheduled_at := scheduled_at sched.
(* Assume that a job cost cannot be larger than a task cost. *)
Hypothesis H_job_cost_le_task_cost:
cost_of_jobs_from_arrival_sequence_le_task_cost
task_cost job_cost job_task arr_seq.
(* We say that the model is fully nonpreemptive
iff every job cannot be preempted until its completion. *)
Definition can_be_preempted_for_fully_nonpreemptive_model (j: Job) (progr: time) :=
(progr == 0) || (progr == job_cost j).
(* Since in a fully nonpreemptive model a job cannot be preempted after
it starts the execution, job_max_nps is equal to job_cost. *)
Let job_max_nps (j: Job) := job_cost j.
(* In order to bound job_max_nps, task_max_nps should be equal to task_cost. *)
Let task_max_nps (tsk: Task) := task_cost tsk.
(* Then, we prove that fully_nonpreemptive_model is a correct preemption model... *)
Lemma fully_nonpreemptive_model_is_correct:
correct_preemption_model arr_seq sched can_be_preempted_for_fully_nonpreemptive_model.
Proof.
intros j; split.
{ move => t.
rewrite /can_be_preempted_for_fully_nonpreemptive_model Bool.negb_orb -lt0n.
move => /andP [POS NCOMPL].
unfold NonpreemptiveSchedule.is_nonpreemptive_schedule in *.
move: (incremental_service_during _ _ _ _ _ POS) => [ft [/andP [_ LT] [SCHED SERV]]].
apply H_nonpreemptive_sched with ft.
{ by apply ltnW. }
{ by done. }
{ rewrite /completed_by neq_ltn; apply/orP; left.
move: NCOMPL; rewrite neq_ltn; move => /orP [LE|GE]; [by done | exfalso].
move: GE; rewrite ltnNge; move => /negP GE; apply: GE.
by eauto 2.
}
}
{ intros t NSCHED SCHED.
rewrite /can_be_preempted_for_fully_nonpreemptive_model.
apply/orP; left.
apply/negP; intros CONTR.
move: CONTR => /negP; rewrite -lt0n; intros POS.
move: (incremental_service_during _ _ _ _ _ POS) => [ft [/andP [_ LT] [SCHEDn SERV]]].
move: NSCHED => /negP NSCHED; apply: NSCHED.
apply H_nonpreemptive_sched with ft.
{ by rewrite -ltnS. }
{ by done. }
{ rewrite /completed_by neq_ltn; apply/orP; left.
apply leq_ltn_trans with (service sched j t.+1).
{ by rewrite /service /service_during big_nat_recr //= leq_addr. }
{ rewrite -addn1.
apply leq_trans with (service sched j t.+2).
- unfold service, service_during.
have EQ: (service_at sched j t.+1) = 1.
{ by apply/eqP; rewrite eqb1. }
by rewrite -EQ -big_nat_recr //=.
- by eauto 2.
}
}
}
Qed.
(* ... and has bounded nonpreemptive regions. *)
Lemma fully_nonpreemptive_model_is_model_with_bounded_nonpreemptive_regions:
model_with_bounded_nonpreemptive_segments
job_cost job_task arr_seq can_be_preempted_for_fully_nonpreemptive_model job_max_nps task_max_nps.
Proof.
have F: forall n, n = 0 \/ n > 0.
{ by intros n; destruct n; [left | right]. }
intros j; split; last split; last split.
{ by done. }
{ by apply/orP; right. }
{ intros ARR.
rewrite /job_max_nps /task_max_nps.
by eauto 2.
}
{ intros progr.
move: (F (progr)) => [EQ | GT].
{ exists progr; split.
- by apply/andP; split; [done | rewrite leq_addr].
- by rewrite /can_be_preempted_for_fully_nonpreemptive_model EQ. }
{ exists (maxn progr (job_cost j)).
have POS: 0 < job_cost j.
{ by apply leq_trans with progr; last move: H0 => /andP [_ H0]. }
split.
{ apply/andP; split; first by rewrite leq_maxl.
rewrite /job_max_nps addnBA; last eauto 2.
rewrite geq_max; apply/andP; split.
- rewrite -addnBA; last by eauto 2.
by rewrite leq_addr.
- by rewrite addnC -addnBA // leq_addr.
}
{ unfold can_be_preempted_for_fully_nonpreemptive_model.
apply/orP; right.
move: H0 => /andP [_ LE].
rewrite eqn_leq; apply/andP; split.
- by rewrite geq_max; apply/andP; split.
- by rewrite leq_max; apply/orP; right.
}
}
}
Qed.
End FullyNonPreemptiveModel.
End FullyNonPreemptivePlatform.
\ No newline at end of file
Require Import rt.util.all.
Require Import rt.model.arrival.basic.job
rt.model.arrival.basic.task
rt.model.priority
rt.model.arrival.basic.task_arrival.
Require Import rt.model.schedule.uni.schedule
rt.model.schedule.uni.service
rt.model.schedule.uni.basic.platform.
Require Export rt.model.schedule.uni.limited.platform.definitions.
From mathcomp Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq fintype bigop.
(** * Platform for fully premptive model *)
(** In module uni.limited.platform we introduce the notion of whether a job can be preempted
at a given time (using a predicate can_be_preempted). In this section, we instantiate
can_be_preempted for the fully preemptive model and prove its correctness. *)
Module FullyPreemptivePlatform.
Import Epsilon Job SporadicTaskset UniprocessorSchedule Priority
Service LimitedPreemptionPlatform.
Section FullyPreemptiveModel.
Context {Task: eqType}.
Context {Job: eqType}.
Variable job_arrival: Job -> time.
Variable job_cost: Job -> time.
Variable job_task: Job -> Task.
(* Consider any job arrival sequence ...*)
Variable arr_seq: arrival_sequence Job.
(* ...and any uniprocessor schedule of these jobs. *)
Variable sched: schedule Job.
(* For simplicity, let's define some local names. *)
Let job_pending := pending job_arrival job_cost sched.
Let job_completed_by := completed_by job_cost sched.
Let job_scheduled_at := scheduled_at sched.
(* In the fully preemptive model any job can be preempted at any time. *)
Definition can_be_preempted_for_fully_preemptive_model (j: Job) (progr: time) := true.
(* Since in a fully preemptive model a job can be preempted at
any time job_max_nps cannot be greater than ε. *)
Let job_max_nps (j: Job) := ε.
(* In order to bound job_max_nps, we can choose task_max_nps that is equal to ε for any task. *)
Let task_max_nps (tsk: Task) := ε.
(* Then, we prove that fully_preemptive_model is a correct preemption model... *)
Lemma fully_preemptive_model_is_correct:
correct_preemption_model arr_seq sched can_be_preempted_for_fully_preemptive_model.
Proof. by intros j; split; intros t CONTR. Qed.
(* ... and has bounded nonpreemptive regions. *)
Lemma fully_preemptive_model_is_model_with_bounded_nonpreemptive_regions:
model_with_bounded_nonpreemptive_segments
job_cost job_task arr_seq can_be_preempted_for_fully_preemptive_model job_max_nps task_max_nps.
Proof.
intros j; repeat split; try done.
intros t; exists t; split.
{ by apply/andP; split; [ done | rewrite subnn addn0]. }
{ by done. }
Qed.
End FullyPreemptiveModel.
End FullyPreemptivePlatform.
\ No newline at end of file
This diff is collapsed.
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