Library prosa.analysis.facts.model.service_of_jobs
Require Export prosa.model.aggregate.workload.
Require Export prosa.model.aggregate.service_of_jobs.
Require Export prosa.analysis.facts.behavior.completion.
Require Export prosa.analysis.facts.model.ideal_schedule.
Require Import prosa.model.processor.ideal.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq fintype bigop.
Require Export prosa.model.aggregate.service_of_jobs.
Require Export prosa.analysis.facts.behavior.completion.
Require Export prosa.analysis.facts.model.ideal_schedule.
Require Import prosa.model.processor.ideal.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq fintype bigop.
Lemmas about Service Received by Sets of Jobs
In this file, we establish basic facts about the service received by sets of jobs.
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 kind of processor state model, ...
... any job arrival sequence with consistent arrivals, ....
Variable arr_seq : arrival_sequence Job.
Hypothesis H_arrival_times_are_consistent : consistent_arrival_times arr_seq.
Hypothesis H_arrival_times_are_consistent : consistent_arrival_times arr_seq.
... and any schedule of this arrival sequence ...
... where jobs do not execute before their arrival or after completion.
Hypothesis H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched.
Hypothesis H_completed_jobs_dont_execute : completed_jobs_dont_execute sched.
Hypothesis H_completed_jobs_dont_execute : completed_jobs_dont_execute sched.
Let P be any predicate over jobs.
Let's define some local names for clarity.
In this section, we prove that the service received by any set of jobs
is upper-bounded by the corresponding workload.
Let jobs denote any (finite) set of jobs.
Assume that the processor model is a unit service model. I.e.,
no job ever receives more than one unit of service at any time.
Then, we prove that the service received by those jobs is no larger than their workload.
Lemma service_of_jobs_le_workload:
∀ t1 t2,
service_of_jobs sched P jobs t1 t2 ≤ workload_of_jobs P jobs.
End ServiceBoundedByWorkload.
∀ t1 t2,
service_of_jobs sched P jobs t1 t2 ≤ workload_of_jobs P jobs.
End ServiceBoundedByWorkload.
In this section we prove a few facts about splitting
the total service of a set of jobs.
We show that the total service of jobs released in a time interval t1,t2)
during [t1,t2) is equal to the sum of:
(1) the total service of jobs released in time interval [t1, t) during time [t1, t)
(2) the total service of jobs released in time interval [t1, t) during time [t, t2)
and (3) the total service of jobs released in time interval [t, t2) during time [t, t2).
Lemma service_of_jobs_cat_scheduling_interval :
∀ t1 t2 t,
t1 ≤ t ≤ t2 →
service_of_jobs sched P (arrivals_between t1 t2) t1 t2
= service_of_jobs sched P (arrivals_between t1 t) t1 t
+ service_of_jobs sched P (arrivals_between t1 t) t t2
+ service_of_jobs sched P (arrivals_between t t2) t t2.
∀ t1 t2 t,
t1 ≤ t ≤ t2 →
service_of_jobs sched P (arrivals_between t1 t2) t1 t2
= service_of_jobs sched P (arrivals_between t1 t) t1 t
+ service_of_jobs sched P (arrivals_between t1 t) t t2
+ service_of_jobs sched P (arrivals_between t t2) t t2.
We show that the total service of jobs released in a time interval t1,t2)
during [t,t2) is equal to the sum of:
(1) the total service of jobs released in a time interval [t1,t) during [t,t2)
and (2) the total service of jobs released in a time interval [t,t2) during [t,t2).
Lemma service_of_jobs_cat_arrival_interval :
∀ t1 t2 t,
t1 ≤ t ≤ t2 →
service_of_jobs sched P (arrivals_between t1 t2) t t2 =
service_of_jobs sched P (arrivals_between t1 t) t t2 +
service_of_jobs sched P (arrivals_between t t2) t t2.
End ServiceCat.
End GenericModelLemmas.
∀ t1 t2 t,
t1 ≤ t ≤ t2 →
service_of_jobs sched P (arrivals_between t1 t2) t t2 =
service_of_jobs sched P (arrivals_between t1 t) t t2 +
service_of_jobs sched P (arrivals_between t t2) t t2.
End ServiceCat.
End GenericModelLemmas.
In this section, we prove some properties about service
of sets of jobs for ideal uni-processor model.
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 arrival sequence with consistent arrivals.
Variable arr_seq : arrival_sequence Job.
Hypothesis H_arrival_times_are_consistent : consistent_arrival_times arr_seq.
Hypothesis H_arrival_times_are_consistent : consistent_arrival_times arr_seq.
Next, consider any ideal uni-processor schedule of this arrival sequence ...
Variable sched : schedule (ideal.processor_state Job).
Hypothesis H_jobs_come_from_arrival_sequence:
jobs_come_from_arrival_sequence sched arr_seq.
Hypothesis H_jobs_come_from_arrival_sequence:
jobs_come_from_arrival_sequence sched arr_seq.
... where jobs do not execute before their arrival or after completion.
Hypothesis H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched.
Hypothesis H_completed_jobs_dont_execute : completed_jobs_dont_execute sched.
Hypothesis H_completed_jobs_dont_execute : completed_jobs_dont_execute sched.
Let P be any predicate over jobs.
Let's define some local names for clarity.
We prove that if the total service within some time interval [t1,t2)]
is smaller than [t2-t1], then there is an idle time instant t ∈ [[t1,t2)].
Lemma low_service_implies_existence_of_idle_time :
∀ t1 t2,
service_of_jobs sched predT (arrivals_between 0 t2) t1 t2 < t2 - t1 →
∃ t, t1 ≤ t < t2 ∧ is_idle sched t.
∀ t1 t2,
service_of_jobs sched predT (arrivals_between 0 t2) t1 t2 < t2 - t1 →
∃ t, t1 ≤ t < t2 ∧ is_idle sched t.
In this section, we prove that the service received by any set of jobs
is upper-bounded by the corresponding interval length.
Let jobs denote any (finite) set of jobs.
Assume that the sequence of jobs is a set.
We prove that the overall service of jobs at a single time instant is at most 1.
Then, we prove that the service received by those jobs is no larger than their workload.
Corollary service_of_jobs_le_length_of_interval:
∀ (t : instant) (Δ : duration),
service_of_jobs sched P jobs t (t + Δ) ≤ Δ.
∀ (t : instant) (Δ : duration),
service_of_jobs sched P jobs t (t + Δ) ≤ Δ.
Similar lemma holds for two instants.
Corollary service_of_jobs_le_length_of_interval':
∀ (t1 t2 : instant),
service_of_jobs sched P jobs t1 t2 ≤ t2 - t1.
End ServiceOfJobsIsBoundedByLength.
∀ (t1 t2 : instant),
service_of_jobs sched P jobs t1 t2 ≤ t2 - t1.
End ServiceOfJobsIsBoundedByLength.
In this section, we introduce a connection between the cumulative
service, cumulative workload, and completion of jobs.
Next, we consider some time instant t_compl.
And state the proposition that all jobs are completed by time
t_compl. Next we show that this proposition is equivalent to
the fact that workload of jobs = service of jobs.
First, we prove that if the workload of jobs is equal to the service
of jobs, then any job in jobs is completed by time t_compl.
Lemma workload_eq_service_impl_all_jobs_have_completed:
workload_of_jobs P jobs =
service_of_jobs sched P jobs t1 t_compl →
all_jobs_completed_by t_compl.
workload_of_jobs P jobs =
service_of_jobs sched P jobs t1 t_compl →
all_jobs_completed_by t_compl.
And vice versa, the fact that any job in jobs is completed by time t_compl
implies that the workload of jobs is equal to the service of jobs.
Lemma all_jobs_have_completed_impl_workload_eq_service:
all_jobs_completed_by t_compl →
workload_of_jobs P jobs =
service_of_jobs sched P jobs t1 t_compl.
all_jobs_completed_by t_compl →
workload_of_jobs P jobs =
service_of_jobs sched P jobs t1 t_compl.
Using the lemmas above, we prove equivalence.