Skip to content
Snippets Groups Projects
Commit 57806e76 authored by Marco Maida's avatar Marco Maida Committed by mmaida
Browse files

Work-conservation transformation and proof.

Added the proof of correctness of the work-conservation transformation for an ideal uniprocessor schedule.
parent c8cff9a3
No related branches found
No related tags found
No related merge requests found
......@@ -374,6 +374,18 @@ Section RelationToScheduled.
rewrite /service_during big_nat_eq0 => IS_ZERO.
by apply (IS_ZERO t); apply /andP; split => //.
Qed.
(** Conversely, if a job is not scheduled during an interval, then
it does not receive any service in that interval *)
Lemma not_scheduled_during_implies_zero_service:
forall t1 t2,
(forall t, t1 <= t < t2 -> ~~ scheduled_at sched j t) ->
service_during sched j t1 t2 = 0.
Proof.
intros t1 t2 NSCHED.
apply big_nat_eq0; move=> t NEQ.
by apply no_service_not_scheduled, NSCHED.
Qed.
(** If a job is scheduled at some point in an interval, it receives
positive cumulative service during the interval... *)
......@@ -578,3 +590,56 @@ Section RelationToScheduled.
End RelationToScheduled.
Section ServiceInTwoSchedules.
(** Consider any job type and any processor model. *)
Context {Job: JobType}.
Context {PState: Type}.
Context `{ProcessorState Job PState}.
(** Consider any two given schedules... *)
Variable sched1 sched2: schedule PState.
(** Given an interval in which the schedules provide the same service
to a job at each instant, we can prove that the cumulative service
received during the interval has to be the same. *)
Section ServiceDuringEquivalentInterval.
(** Consider two time instants... *)
Variable t1 t2 : instant.
(** ...and a given job that is to be scheduled. *)
Variable j: Job.
(** Assume that, in any instant between [t1] and [t2] the service
provided to [j] from the two schedules is the same. *)
Hypothesis H_sched1_sched2_same_service_at:
forall t, t1 <= t < t2 ->
service_at sched1 j t = service_at sched2 j t.
(** It follows that the service provided during [t1] and [t2]
is also the same. *)
Lemma same_service_during:
service_during sched1 j t1 t2 = service_during sched2 j t1 t2.
Proof.
rewrite /service_during.
apply eq_big_nat.
by apply H_sched1_sched2_same_service_at.
Qed.
End ServiceDuringEquivalentInterval.
(** We can leverage the previous lemma to conclude that two schedules
that match in a given interval will also have the same cumulative
service across the interval. *)
Corollary equal_prefix_implies_same_service_during:
forall t1 t2,
(forall t, t1 <= t < t2 -> sched1 t = sched2 t) ->
forall j, service_during sched1 j t1 t2 = service_during sched2 j t1 t2.
Proof.
move=> t1 t2 SCHED_EQ j.
apply same_service_during => t' RANGE.
by rewrite /service_at SCHED_EQ.
Qed.
End ServiceInTwoSchedules.
This diff is collapsed.
Require Export prosa.analysis.transform.swap.
Require Export prosa.analysis.transform.prefix.
Require Export prosa.util.search_arg.
Require Export prosa.util.list.
(** Throughout this file, we assume ideal uniprocessor schedules. *)
Require Export prosa.model.processor.ideal.
(** In this file we define the transformation from any ideal uniprocessor schedule
into a work-conserving one. The procedure is to patch the idle allocations
with future job allocations. Note that a job cannot be allocated before
its arrival, therefore there could still exist idle instants between any two
job allocations. *)
Section WCTransformation.
(** Consider any type of jobs with arrival times, costs, and deadlines... *)
Context {Job : JobType}.
Context `{JobArrival Job}.
Context `{JobCost Job}.
Context `{JobDeadline Job}.
(** ...an ideal uniprocessor schedule... *)
Let PState := ideal.processor_state Job.
(** ...and any valid job arrival sequence. *)
Variable arr_seq : arrival_sequence Job.
Hypothesis H_arr_seq_valid : valid_arrival_sequence arr_seq.
(** We say that a state is relevant (for the purpose of the
transformation) if it is not idle and if the job scheduled in it
has arrived prior to some given reference time. *)
Definition relevant_pstate reference_time pstate :=
match pstate with
| None => false
| Some j => job_arrival j <= reference_time
end.
(** In order to patch an idle allocation, we look in the future for another allocation
that could be moved there. The limit of the search is the maximum deadline of
every job arrived before the given moment. *)
Definition max_deadline_for_jobs_arrived_before arrived_before :=
let deadlines := map job_deadline (arrivals_up_to arr_seq arrived_before)
in max0 deadlines.
(** Next, we define a central element of the work-conserving transformation
procedure: given an idle allocation at [t], find a job allocation in the future
to swap with. *)
Definition find_swap_candidate sched t :=
let order _ _ := false (* always take the first result *)
in
let max_dl := max_deadline_for_jobs_arrived_before t
in
let search_result := search_arg sched (relevant_pstate t) order t max_dl
in
if search_result is Some t_swap
then t_swap
else t. (* if nothing is found, swap with yourself *)
(** The point-wise transformation procedure: given a schedule and a
time [t1], ensure that the schedule is work-conserving at time
[t1]. *)
Definition make_wc_at sched t1 :=
match sched t1 with
| Some j => sched (* leave working instants alone *)
| None =>
let
t2 := find_swap_candidate sched t1
in swapped sched t1 t2
end.
(** To transform a finite prefix of a given reference schedule, apply
[make_wc_at] to every point up to the given finite horizon. *)
Definition wc_transform_prefix sched horizon :=
prefix_map sched make_wc_at horizon.
(** Finally, a fully work-conserving schedule (i.e., one that is
work-conserving at any time) is obtained by first computing a
work-conserving prefix up to and including the requested time [t],
and by then looking at the last point of the prefix. *)
Definition wc_transform sched t :=
let
wc_prefix := wc_transform_prefix sched t.+1
in wc_prefix t.
End WCTransformation.
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