Library prosa.analysis.facts.model.task_schedule
Require Export prosa.model.task.concept.
Require Export prosa.analysis.definitions.task_schedule.
Require Export prosa.analysis.facts.behavior.service.
Require Export prosa.analysis.facts.model.scheduled.
Require Export prosa.analysis.definitions.task_schedule.
Require Export prosa.analysis.facts.behavior.service.
Require Export prosa.analysis.facts.model.scheduled.
In this file we provide basic properties related to schedule of a task.
Consider any type of tasks ...
... and any type of jobs associated with these tasks.
Context {Job : JobType}.
Context `{JobTask Job Task}.
Context `{JobArrival Job}.
Context `{JobCost Job}.
Context `{JobTask Job Task}.
Context `{JobArrival Job}.
Context `{JobCost Job}.
Consider any valid arrival sequence of such jobs ...
Variable arr_seq : arrival_sequence Job.
Hypothesis H_valid_arrivals : valid_arrival_sequence arr_seq.
Hypothesis H_valid_arrivals : valid_arrival_sequence arr_seq.
... and let sched be any corresponding uni-processor schedule.
Context {PState : ProcessorState Job}.
Hypothesis H_uniproc : uniprocessor_model PState.
Variable sched : schedule PState.
Hypothesis H_jobs_come_from_arrival_sequence :
jobs_come_from_arrival_sequence sched arr_seq.
Hypothesis H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched.
Hypothesis H_uniproc : uniprocessor_model PState.
Variable sched : schedule PState.
Hypothesis H_jobs_come_from_arrival_sequence :
jobs_come_from_arrival_sequence sched arr_seq.
Hypothesis H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched.
Let tsk be any task.
Lemma task_served_task_scheduled :
∀ t,
task_served_at arr_seq sched tsk t →
task_scheduled_at arr_seq sched tsk t.
∀ t,
task_served_at arr_seq sched tsk t →
task_scheduled_at arr_seq sched tsk t.
Next, we show that, under ideal-progress assumption, the notion
of task served coincides with the notion of task scheduled.
Lemma task_served_eq_task_scheduled :
ideal_progress_proc_model PState →
∀ t,
task_served_at arr_seq sched tsk t = task_scheduled_at arr_seq sched tsk t.
ideal_progress_proc_model PState →
∀ t,
task_served_at arr_seq sched tsk t = task_scheduled_at arr_seq sched tsk t.
We note that if the processor is idle at time t, then no task
is scheduled.
Lemma no_task_scheduled_when_idle :
∀ t,
is_idle arr_seq sched t →
~~ task_scheduled_at arr_seq sched tsk t.
∀ t,
is_idle arr_seq sched t →
~~ task_scheduled_at arr_seq sched tsk t.
Similarly, if the processor is idle at time t, then no task is
served.
Lemma no_task_served_when_idle :
∀ t,
is_idle arr_seq sched t →
~~ task_served_at arr_seq sched tsk t.
∀ t,
is_idle arr_seq sched t →
~~ task_served_at arr_seq sched tsk t.
We show that if a job is scheduled at time t, then
task_scheduled_at tsk t is equal to job_of_task tsk j. In
other words, any occurrence of task_scheduled_at can be
replaced with job_of_task.
Lemma job_of_scheduled_task :
∀ j t,
scheduled_at sched j t →
task_scheduled_at arr_seq sched tsk t = job_of_task tsk j.
∀ j t,
scheduled_at sched j t →
task_scheduled_at arr_seq sched tsk t = job_of_task tsk j.
As a corollary, we show that if a job of task tsk is scheduled
at time t, then task tsk is scheduled at time t.
Corollary job_of_task_scheduled :
∀ j t,
job_of_task tsk j →
scheduled_at sched j t →
task_scheduled_at arr_seq sched tsk t.
∀ j t,
job_of_task tsk j →
scheduled_at sched j t →
task_scheduled_at arr_seq sched tsk t.
And vice versa, if no job of task tsk is scheduled at time
t, then task tsk is not scheduled at time t.
Corollary job_of_other_task_scheduled :
∀ j t,
~~ job_of_task tsk j →
scheduled_at sched j t →
~~ task_scheduled_at arr_seq sched tsk t.
∀ j t,
~~ job_of_task tsk j →
scheduled_at sched j t →
~~ task_scheduled_at arr_seq sched tsk t.
Similarly, we show that if no job of task tsk is scheduled at
time t, then task tsk is not served at time t.
Corollary job_of_other_task_scheduled' :
∀ j t,
~~ job_of_task tsk j →
scheduled_at sched j t →
~~ task_served_at arr_seq sched tsk t.
∀ j t,
~~ job_of_task tsk j →
scheduled_at sched j t →
~~ task_served_at arr_seq sched tsk t.
Lastly, if a job of task tsk is scheduled at time t but
receives no service, then task tsk is not served at time
t.
Corollary job_of_task_not_served :
∀ j t,
scheduled_at sched j t →
job_of_task tsk j →
service_at sched j t = 0 →
~~ task_served_at arr_seq sched tsk t.
∀ j t,
scheduled_at sched j t →
job_of_task tsk j →
service_at sched j t = 0 →
~~ task_served_at arr_seq sched tsk t.
In the following section, we prove a rewriting rule between the
predicate task_served_at and job_of_task.
Assume that the processor is fully supply-consuming.
Consider a time instant t ...
... and assume that there is supply at t.
Consider jobs j and j' (j' is not necessarily distinct
from job j). Assume that j is scheduled at time t.
Then the predicate task_served_at is equal to the predicate
job_of_task.
Lemma task_served_at_eq_job_of_task :
task_served_at arr_seq sched tsk t = job_of_task tsk j.
End SomeJobIsScheduled.
End TaskSchedule.
task_served_at arr_seq sched tsk t = job_of_task tsk j.
End SomeJobIsScheduled.
End TaskSchedule.