Library prosa.classic.model.schedule.uni.limited.busy_interval
Require Import prosa.classic.util.all.
Require Import prosa.classic.model.arrival.basic.task
prosa.classic.model.arrival.basic.job
prosa.classic.model.arrival.basic.arrival_sequence
prosa.classic.model.priority
prosa.classic.model.arrival.basic.task_arrival.
Require Import prosa.classic.model.schedule.uni.service
prosa.classic.model.schedule.uni.workload
prosa.classic.model.schedule.uni.schedule.
Require Import prosa.classic.model.schedule.uni.limited.platform.definitions.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq fintype bigop.
Require Import prosa.classic.model.arrival.basic.task
prosa.classic.model.arrival.basic.job
prosa.classic.model.arrival.basic.arrival_sequence
prosa.classic.model.priority
prosa.classic.model.arrival.basic.task_arrival.
Require Import prosa.classic.model.schedule.uni.service
prosa.classic.model.schedule.uni.workload
prosa.classic.model.schedule.uni.schedule.
Require Import prosa.classic.model.schedule.uni.limited.platform.definitions.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq fintype bigop.
Busy Interval for JLFP-models
In this module we define the notion of busy intervals for uniprocessor for JLFP schedulers.
Module BusyIntervalJLFP.
Import Job Priority UniprocessorSchedule LimitedPreemptionPlatform Service Workload TaskArrival.
Section Definitions.
Context {Task: eqType}.
Context {Job: eqType}.
Variable job_arrival: Job → time.
Variable job_cost: Job → time.
Variable job_task: Job → Task.
(* Consider any arrival sequence with consistent arrival times... *)
Variable arr_seq: arrival_sequence Job.
Hypothesis H_arrival_times_are_consistent: arrival_times_are_consistent job_arrival arr_seq.
(* ...and any uniprocessor schedule of these jobs. *)
Variable sched: schedule Job.
Hypothesis H_jobs_come_from_arrival_sequence:
jobs_come_from_arrival_sequence sched arr_seq.
(* Assume a given JLFP policy. *)
Variable higher_eq_priority: JLFP_policy Job.
(* For simplicity, let's define some local names. *)
Let job_pending_at := pending job_arrival job_cost sched.
Let job_scheduled_at := scheduled_at sched.
Let job_completed_by := completed_by job_cost sched.
Let job_remaining_cost j t := remaining_cost job_cost sched j t.
Let arrivals_between := jobs_arrived_between arr_seq.
(* In this section, we define the notion of a busy interval. *)
Section BusyInterval.
(* Consider an arbitrary task tsk. *)
Variable tsk: Task.
(* Consider any job j of tsk. *)
Variable j: Job.
Hypothesis H_from_arrival_sequence: arrives_in arr_seq j.
Hypothesis H_job_task: job_task j = tsk.
(* We say that t is a quiet time for j iff every higher-priority job from
the arrival sequence that arrived before t has completed by that time. *)
Definition quiet_time (t: time) :=
∀ j_hp,
arrives_in arr_seq j_hp →
higher_eq_priority j_hp j →
arrived_before job_arrival j_hp t →
job_completed_by j_hp t.
(* Based on the definition of quiet time, we say that interval
t1, t_busy) is a (potentially unbounded) busy-interval prefix iff the interval starts with a quiet time where a higher or equal priority job is released and remains non-quiet. We also require job j to be release in the interval. *)
Definition busy_interval_prefix (t1 t_busy: time) :=
t1 < t_busy ∧
quiet_time t1 ∧
(∀ t, t1 < t < t_busy → ¬ quiet_time t) ∧
t1 ≤ job_arrival j < t_busy.
(* Next, we say that an interval t1, t2) is a busy interval iff [t1, t2) is a busy-interval prefix and t2 is a quiet time. *)
Definition busy_interval (t1 t2: time) :=
busy_interval_prefix t1 t2 ∧
quiet_time t2.
End BusyInterval.
(* In this section, we define a notion of bounded priority inversion experienced by a job. *)
Section JobPriorityInversionBound.
(* Consider an arbitrary task tsk. *)
Variable tsk: Task.
(* Consider any job j of tsk. *)
Variable j: Job.
Hypothesis H_from_arrival_sequence: arrives_in arr_seq j.
Hypothesis H_job_task: job_task j = tsk.
(* We say that the job incurs priority inversion if it has higher priority than the scheduled
job. Note that this definition implicitly assumes that the scheduler is work-conserving in
the sense of the definition given in prosa.classic.model.schedule.uni.basic.platform. Therefore, it
cannot be applied to models with jitter or self-suspensions. *)
Definition is_priority_inversion t :=
if sched t is Some jlp then
~~ higher_eq_priority jlp j
else false.
(* Then we compute the cumulative priority inversion incurred by
a job within some time interval t1, t2). *)
Definition cumulative_priority_inversion t1 t2 :=
\sum_(t1 ≤ t < t2) is_priority_inversion t.
(* We say that priority inversion of job j is bounded by a constant B iff cumulative
priority inversion within any busy inverval prefix is bounded by B. *)
Definition priority_inversion_of_job_is_bounded_by (B: time) :=
∀ (t1 t2: time),
busy_interval_prefix j t1 t2 →
cumulative_priority_inversion t1 t2 ≤ B.
End JobPriorityInversionBound.
(* In this section, we define a notion of the bounded priority inversion for task. *)
Section TaskPriorityInversionBound.
(* Consider an arbitrary task tsk. *)
Variable tsk: Task.
(* We say that task tsk has bounded priority inversion if all
its jobs have bounded cumulative priority inversion. *)
Definition priority_inversion_is_bounded_by (B: time) :=
∀ (j: Job),
arrives_in arr_seq j →
job_task j = tsk →
job_cost j > 0 →
priority_inversion_of_job_is_bounded_by j B.
End TaskPriorityInversionBound.
(* In this section we define the computational
version of the notion of quiet time. *)
Section DecidableQuietTime.
(* We say that t is a quiet time for j iff every higher-priority job from
the arrival sequence that arrived before t has completed by that time. *)
Definition quiet_time_dec (j : Job) (t : time) :=
all
(fun j_hp ⇒ higher_eq_priority j_hp j ==> (completed_by job_cost sched j_hp t))
(jobs_arrived_before arr_seq t).
(* We also show that the computational and propositional definitions are equivalent. *)
Lemma quiet_time_P :
∀ j t, reflect (quiet_time j t) (quiet_time_dec j t).
End DecidableQuietTime.
(* Now we prove some lemmas about busy intervals. *)
Section Lemmas.
(* Consider an arbitrary task tsk. *)
Variable tsk: Task.
(* Consider an arbitrary job j. *)
Variable j: Job.
Hypothesis H_from_arrival_sequence: arrives_in arr_seq j.
Hypothesis H_job_task: job_task j = tsk.
Hypothesis H_job_cost_positive: job_cost_positive job_cost j.
(* Recall the list of jobs that arrive in any interval. *)
Let quiet_time t1 := quiet_time j t1.
Let busy_interval_prefix t1 t2 := busy_interval_prefix j t1 t2.
Let busy_interval t1 t2 := busy_interval j t1 t2.
Let is_priority_inversion_bounded_by K := priority_inversion_of_job_is_bounded_by j K.
(* We begin by proving a basic lemma about completion of the job within its busy interval. *)
Section BasicLemma.
(* Assume that the priority relation is reflexive. *)
Hypothesis H_priority_is_reflexive: FP_is_reflexive higher_eq_priority.
(* Consider any busy interval t1, t2) of job j. *)
Variable t1 t2: time.
Hypothesis H_busy_interval: busy_interval t1 t2.
(* We prove that job j completes by the end of the busy interval. *)
Lemma job_completes_within_busy_interval:
job_completed_by j t2.
End BasicLemma.
(* In this section, we prove that during a busy interval there
always exists a pending job. *)
Section ExistsPendingJob.
(* Assume that jobs do not execute after completion. *)
Hypothesis H_completed_jobs_dont_execute:
completed_jobs_dont_execute job_cost sched.
(* Let t1, t2 be any interval where time t1 is quiet and time t2 is not quiet. *)
Variable t1 t2: time.
Hypothesis H_interval: t1 ≤ t2.
Hypothesis H_quiet: quiet_time t1.
Hypothesis H_not_quiet: ¬ quiet_time t2.
(* Then, we prove that there is a job pending at time t2
that has higher or equal priority (with respect of tsk). *)
Lemma not_quiet_implies_exists_pending_job:
∃ j_hp,
arrives_in arr_seq j_hp ∧
arrived_between job_arrival j_hp t1 t2 ∧
higher_eq_priority j_hp j ∧
¬ job_completed_by j_hp t2.
End ExistsPendingJob.
(* In this section, we prove that during a busy interval the
processor is never idle. *)
Section ProcessorAlwaysBusy.
(* Assume that the schedule is work-conserving and that jobs do
not execute before their arrival nor after completion. *)
Hypothesis H_work_conserving: work_conserving job_arrival job_cost arr_seq sched.
Hypothesis H_completed_jobs_dont_execute: completed_jobs_dont_execute job_cost sched.
Hypothesis H_jobs_must_arrive_to_execute: jobs_must_arrive_to_execute job_arrival sched.
(* Next, we assume that the priority relation is reflexive and transitive. *)
Hypothesis H_priority_is_reflexive: JLFP_is_reflexive higher_eq_priority.
Hypothesis H_priority_is_transitive: JLFP_is_transitive higher_eq_priority.
(* Consider any busy interval prefix t1, t2). *)
Variable t1 t2: time.
Hypothesis H_busy_interval_prefix: busy_interval_prefix t1 t2.
(* We prove that if the processot is idle at a time instant t, then
the next time instant t+1 will be a quiet time. *)
Lemma idle_time_implies_quiet_time_at_the_next_time_instant:
∀ t,
is_idle sched t →
quiet_time t.+1.
(* Next, we prove that at any time instant t within the busy interval there exists a job
jhp such that (1) job jhp is pending at time t and (2) job jhp has higher-or-equal
priority than task tsk. *)
Lemma pending_hp_job_exists:
∀ t,
t1 ≤ t < t2 →
∃ jhp,
arrives_in arr_seq jhp ∧
job_pending_at jhp t ∧
higher_eq_priority jhp j.
(* We prove that at any time instant t within t1, t2) the processor is not idle. *)
Lemma not_quiet_implies_not_idle:
∀ t,
t1 ≤ t < t2 →
¬ is_idle sched t.
End ProcessorAlwaysBusy.
(* In section we prove a few auxiliary lemmas about quiet time and service. *)
Section QuietTimeAndServiceOfJobs.
(* Assume that there are no duplicate job arrivals... *)
Hypothesis H_arrival_sequence_is_a_set:
arrival_sequence_is_a_set arr_seq.
(* ...and that 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.
(* We also assume that the schedule is work-conserving. *)
Hypothesis H_work_conserving: work_conserving job_arrival job_cost arr_seq sched.
(* Let t1 be a quiet time. *)
Variable t1: time.
Hypothesis H_quiet_time: quiet_time t1.
(* Assume that there is no quiet time in the interval (t1, t1 + Δ]. *)
Variable Δ: time.
Hypothesis H_no_quiet_time: ∀ t, t1 < t ≤ t1 + Δ → ¬ quiet_time t.
(* For clarity, we introduce a notion of the total service of jobs released in
time interval t_beg, t_end) during the time interval [t1, t1 + Δ). *)
Let service_received_by_hep_jobs_released_during t_beg t_end :=
service_of_higher_or_equal_priority_jobs
sched (arrivals_between t_beg t_end) higher_eq_priority j t1 (t1 + Δ).
(* We prove that jobs with higher-than-or-equal priority that
released before time instant t1 receive no service after
time instant t1. *)
Lemma hep_jobs_receive_no_service_before_quiet_time:
service_received_by_hep_jobs_released_during t1 (t1 + Δ) =
service_received_by_hep_jobs_released_during 0 (t1 + Δ).
(* Next we prove that the total service within a "non-quiet"
time interval t1, t1 + Δ) is exactly Δ. *)
Lemma no_idle_time_within_non_quiet_time_interval:
service_of_jobs sched (arrivals_between 0 (t1 + Δ)) predT t1 (t1 + Δ) = Δ.
End QuietTimeAndServiceOfJobs.
(* In this section, we show that the length of any busy interval
is bounded, as long as there is enough supply to accomodate
the workload of tasks with higher or equal priority. *)
Section BoundingBusyInterval.
(* Assume that there are no duplicate job arrivals... *)
Hypothesis H_arrival_sequence_is_a_set:
arrival_sequence_is_a_set arr_seq.
(* ...and that 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.
(* Also assume a work-conserving JLFP schedule, ... *)
Hypothesis H_work_conserving: work_conserving job_arrival job_cost arr_seq sched.
(* ...in which the priority relation is reflexive and transitive. *)
Hypothesis H_priority_is_reflexive: JLFP_is_reflexive higher_eq_priority.
Hypothesis H_priority_is_transitive: JLFP_is_transitive higher_eq_priority.
(* Next, we recall the notion of workload of all jobs released in a given interval
t1, t2) that have higher-or-equal priority w.r.t the job j being analyzed. *)
Let hp_workload t1 t2 :=
workload_of_higher_or_equal_priority_jobs
job_cost (arrivals_between t1 t2) higher_eq_priority j.
(* With regard to the jobs with higher-or-equal priority that are released
in a given interval t1, t2), we also recall the service received by these jobs in the same interval [t1, t2). *)
Let hp_service t1 t2 :=
service_of_higher_or_equal_priority_jobs
sched (arrivals_between t1 t2) higher_eq_priority j t1 t2.
(* Now we begin the proof. First, we show that the busy interval is bounded. *)
Section BoundingBusyInterval.
(* Suppose that job j is pending at time t_busy. *)
Variable t_busy: time.
Hypothesis H_j_is_pending: job_pending_at j t_busy.
(* First, we show that there must exist a busy interval prefix. *)
Section LowerBound.
(* Since job j is pending, there is a (potentially unbounded)
busy interval that starts no later than with the arrival of j. *)
Lemma exists_busy_interval_prefix:
∃ t1,
busy_interval_prefix t1 t_busy.+1 ∧
t1 ≤ job_arrival j ≤ t_busy.
 
End LowerBound.
(* Next we prove that, if there is a point where the requested workload
is upper-bounded by the supply, then the busy interval eventually ends. *)
Section UpperBound.
(* Consider any busy interval prefix of job j. *)
Variable t1: time.
Hypothesis H_is_busy_prefix: busy_interval_prefix t1 t_busy.+1.
(* Let priority_inversion_bound be a constant which bounds
length of a priority inversion. *)
Variable priority_inversion_bound: time.
Hypothesis H_priority_inversion_is_bounded:
is_priority_inversion_bounded_by priority_inversion_bound.
(* Next, assume that for some positive delta, the sum of requested workload
at time t1 + delta and constant priority_inversion_bound is bounded by
delta (i.e., the supply). *)
Variable delta: time.
Hypothesis H_delta_positive: delta > 0.
Hypothesis H_workload_is_bounded:
priority_inversion_bound + hp_workload t1 (t1 + delta) ≤ delta.
(* If there is a quiet time by time (t1 + delta), it trivially follows that
the busy interval is bounded.
Thus, let's consider first the harder case where there is no quiet time,
which turns out to be impossible. *)
Section CannotBeBusyForSoLong.
(* Assume that there is no quiet time in the interval (t1, t1 + delta]. *)
Hypothesis H_no_quiet_time:
∀ t, t1 < t ≤ t1 + delta → ¬ quiet_time t.
(* Since the interval is always non-quiet, the processor is always busy
with tasks of higher-or-equal priority or some lower priority job which was scheduled,
i.e., the sum of service done by jobs with actual arrival time in t1, t1 + delta) and priority inversion equals delta. *)
Lemma busy_interval_has_uninterrupted_service:
delta ≤ priority_inversion_bound + hp_service t1 (t1 + delta).
  
(* Moreover, the fact that the interval is not quiet also implies
that there's more workload requested than service received. *)
Lemma busy_interval_too_much_workload:
hp_workload t1 (t1 + delta) > hp_service t1 (t1 + delta).
(* Using the two lemmas above, we infer that the workload is larger than the
interval length. However, this contradicts the assumption H_workload_is_bounded. *)
Corollary busy_interval_workload_larger_than_interval:
priority_inversion_bound + hp_workload t1 (t1 + delta) > delta.
End CannotBeBusyForSoLong.
(* Since the interval cannot remain busy for so long, we prove that
the busy interval finishes at some point t2 <= t1 + delta. *)
Lemma busy_interval_is_bounded:
∃ t2,
t2 ≤ t1 + delta ∧
busy_interval t1 t2.
  
End UpperBound.
End BoundingBusyInterval.
(* In this section, we show that from a workload bound we can infer
the existence of a busy interval. *)
Section BusyIntervalFromWorkloadBound.
(* Let priority_inversion_bound be a constant that bounds the length of a priority inversion. *)
Variable priority_inversion_bound: time.
Hypothesis H_priority_inversion_is_bounded:
is_priority_inversion_bounded_by priority_inversion_bound.
(* Assume that for some positive delta, the sum of requested workload at
time (t1 + delta) and priority inversion is bounded by delta (i.e., the supply). *)
Variable delta: time.
Hypothesis H_delta_positive: delta > 0.
Hypothesis H_workload_is_bounded:
∀ t, priority_inversion_bound + hp_workload t (t + delta) ≤ delta.
(* Next, we assume that job j has positive cost, from which we can
infer that there is a time in which j is pending. *)
Hypothesis H_positive_cost: job_cost j > 0.
(* Therefore there must exists a busy interval t1, t2) that contains the arrival time of j. *)
Corollary exists_busy_interval:
∃ t1 t2,
t1 ≤ job_arrival j < t2 ∧
t2 ≤ t1 + delta ∧
busy_interval t1 t2.
 
End BusyIntervalFromWorkloadBound.
(* If we know that the workload is bounded, we can also use the
busy interval to infer a response-time bound. *)
Section ResponseTimeBoundFromBusyInterval.
(* Let priority_inversion_bound be a constant that bounds the length of a priority inversion. *)
Variable priority_inversion_bound: time.
Hypothesis H_priority_inversion_is_bounded:
is_priority_inversion_bounded_by priority_inversion_bound.
(* Assume that for some positive delta, the sum of requested workload at
time (t1 + delta) and priority inversion is bounded by delta (i.e., the supply). *)
Variable delta: time.
Hypothesis H_delta_positive: delta > 0.
Hypothesis H_workload_is_bounded:
∀ t, priority_inversion_bound + hp_workload t (t + delta) ≤ delta.
(* Then, job j must complete by (job_arrival j + delta). *)
Lemma busy_interval_bounds_response_time:
job_completed_by j (job_arrival j + delta).
End ResponseTimeBoundFromBusyInterval.
End BoundingBusyInterval.
End Lemmas.
(* In this section, we derive an alternative condition for the existence of
a busy interval. The new condition requires the total workload (instead
of the high-priority workload) generated by the task set to be bounded. *)
Section NonOverloadedProcessor.
(* The processor has no carry-in at time t iff every job (regardless of priority)
from the arrival sequence released before t has completed by that time. *)
Definition no_carry_in (t: time) :=
∀ j_o,
arrives_in arr_seq j_o →
arrived_before job_arrival j_o t →
job_completed_by j_o t.
(* The fact that there is no carry-in at time instant t
trivially implies that t is a quiet time. *)
Lemma no_carry_in_implies_quiet_time :
∀ j t,
no_carry_in t →
quiet_time j t.
(* Assume that there are no duplicate job arrivals. *)
Hypothesis H_arrival_sequence_is_a_set:
arrival_sequence_is_a_set arr_seq.
(* We also assume that the schedule is work-conserving and that jobs
do not execute before their arrival nor after completion. *)
Hypothesis H_work_conserving: work_conserving job_arrival job_cost arr_seq sched.
Hypothesis H_completed_jobs_dont_execute: completed_jobs_dont_execute job_cost sched.
Hypothesis H_jobs_must_arrive_to_execute: jobs_must_arrive_to_execute job_arrival sched.
(* Next we show that an idle time implies no carry in at this time instant. *)
Lemma idle_instant_implies_no_carry_in_at_t :
∀ t,
is_idle sched t →
no_carry_in t.
(* Moreover, an idle time implies no carry in at the next time instant. *)
Lemma idle_instant_implies_no_carry_in_at_t_pl_1 :
∀ t,
is_idle sched t →
no_carry_in t.+1.
(* Let the priority relation be reflexive. *)
Hypothesis H_priority_is_reflexive: JLFP_is_reflexive higher_eq_priority.
(* Next, we recall the notion of workload of all jobs released in a given interval t1, t2)... *)
Let total_workload t1 t2 :=
workload_of_jobs job_cost (arrivals_between t1 t2) predT.
(* ... and total service of jobs within some time interval t1, t2). *)
Let total_service t1 t2 :=
service_of_jobs sched (arrivals_between 0 t2) predT t1 t2.
(* Assume that for some positive Δ, the sum of requested workload
at time (t + Δ) is bounded by delta (i.e., the supply).
Note that this assumption bounds the total workload of
jobs released in a time interval t, t + Δ) regardless of their priorities. *)
Variable Δ: time.
Hypothesis H_delta_positive: Δ > 0.
Hypothesis H_workload_is_bounded: ∀ t, total_workload t (t + Δ) ≤ Δ.
(* Next we prove that since for any time instant t there is a point where
the total workload is upper-bounded by the supply the processor encounters
no carry-in instants at least every Δ time units. *)
Section ProcessorIsNotTooBusy.
(* We start by proving that the processor has no carry-in at
the beginning (i.e., has no carry-in at time instant 0). *)
Lemma no_carry_in_at_the_beginning :
no_carry_in 0.
(* In this section, we prove that for any time instant t there
exists another time instant t' ∈ (t, t + Δ] such that the
processor has no carry-in at time t'. *)
Section ProcessorIsNotTooBusyInduction.
(* Consider an arbitrary time instant t... *)
Variable t: time.
(* ...such that the processor has no carry-in at time t. *)
Hypothesis H_no_carry_in: no_carry_in t.
(* First, recall that the total service is bounded by the total workload. Therefore
the total service of jobs in the interval t, t + Δ) is bounded by Δ. *)
Lemma total_service_is_bounded_by_Δ :
total_service t (t + Δ) ≤ Δ.
(* Next we consider two cases:
(1) The case when the total service is strictly less than Δ,
(2) And when the total service is equal to Δ. *)
(* In the first case, we use the pigeonhole principle to conclude
that there is an idle time instant; which in turn implies existence
of a time instant with no carry-in. *)
Lemma low_total_service_implies_existence_of_time_with_no_carry_in :
total_service t (t + Δ) < Δ →
∃ δ, δ < Δ ∧ no_carry_in (t.+1 + δ).
(* In the second case, the total service within the time interval t, t + Δ) is equal to Δ. On the other hand, we know that the total workload is lower-bounded by the total service and upper-bounded by Δ. Therefore, the total workload is equal to total service this implies completion of all jobs by time [t + Δ] and hence no carry-in at time [t + Δ]. *)
Lemma completion_of_all_jobs_implies_no_carry_in :
total_service t (t + Δ) = Δ →
no_carry_in (t + Δ).
End ProcessorIsNotTooBusyInduction.
(* Finally, we show that any interval of length Δ contains a time instant with no carry-in. *)
Lemma processor_is_not_too_busy :
∀ t, ∃ δ, δ < Δ ∧ no_carry_in (t + δ).
 
End ProcessorIsNotTooBusy.
(* Consider an arbitrary job j with positive cost. *)
Variable j: Job.
Hypothesis H_from_arrival_sequence: arrives_in arr_seq j.
Hypothesis H_job_cost_positive: job_cost_positive job_cost j.
(* We show that there must exists a busy interval t1, t2) that contains the arrival time of j. *)
Corollary exists_busy_interval_from_total_workload_bound :
∃ t1 t2,
t1 ≤ job_arrival j < t2 ∧
t2 ≤ t1 + Δ ∧
busy_interval j t1 t2.
 
End NonOverloadedProcessor.
End Definitions.
End BusyIntervalJLFP.
Import Job Priority UniprocessorSchedule LimitedPreemptionPlatform Service Workload TaskArrival.
Section Definitions.
Context {Task: eqType}.
Context {Job: eqType}.
Variable job_arrival: Job → time.
Variable job_cost: Job → time.
Variable job_task: Job → Task.
(* Consider any arrival sequence with consistent arrival times... *)
Variable arr_seq: arrival_sequence Job.
Hypothesis H_arrival_times_are_consistent: arrival_times_are_consistent job_arrival arr_seq.
(* ...and any uniprocessor schedule of these jobs. *)
Variable sched: schedule Job.
Hypothesis H_jobs_come_from_arrival_sequence:
jobs_come_from_arrival_sequence sched arr_seq.
(* Assume a given JLFP policy. *)
Variable higher_eq_priority: JLFP_policy Job.
(* For simplicity, let's define some local names. *)
Let job_pending_at := pending job_arrival job_cost sched.
Let job_scheduled_at := scheduled_at sched.
Let job_completed_by := completed_by job_cost sched.
Let job_remaining_cost j t := remaining_cost job_cost sched j t.
Let arrivals_between := jobs_arrived_between arr_seq.
(* In this section, we define the notion of a busy interval. *)
Section BusyInterval.
(* Consider an arbitrary task tsk. *)
Variable tsk: Task.
(* Consider any job j of tsk. *)
Variable j: Job.
Hypothesis H_from_arrival_sequence: arrives_in arr_seq j.
Hypothesis H_job_task: job_task j = tsk.
(* We say that t is a quiet time for j iff every higher-priority job from
the arrival sequence that arrived before t has completed by that time. *)
Definition quiet_time (t: time) :=
∀ j_hp,
arrives_in arr_seq j_hp →
higher_eq_priority j_hp j →
arrived_before job_arrival j_hp t →
job_completed_by j_hp t.
(* Based on the definition of quiet time, we say that interval
t1, t_busy) is a (potentially unbounded) busy-interval prefix iff the interval starts with a quiet time where a higher or equal priority job is released and remains non-quiet. We also require job j to be release in the interval. *)
Definition busy_interval_prefix (t1 t_busy: time) :=
t1 < t_busy ∧
quiet_time t1 ∧
(∀ t, t1 < t < t_busy → ¬ quiet_time t) ∧
t1 ≤ job_arrival j < t_busy.
(* Next, we say that an interval t1, t2) is a busy interval iff [t1, t2) is a busy-interval prefix and t2 is a quiet time. *)
Definition busy_interval (t1 t2: time) :=
busy_interval_prefix t1 t2 ∧
quiet_time t2.
End BusyInterval.
(* In this section, we define a notion of bounded priority inversion experienced by a job. *)
Section JobPriorityInversionBound.
(* Consider an arbitrary task tsk. *)
Variable tsk: Task.
(* Consider any job j of tsk. *)
Variable j: Job.
Hypothesis H_from_arrival_sequence: arrives_in arr_seq j.
Hypothesis H_job_task: job_task j = tsk.
(* We say that the job incurs priority inversion if it has higher priority than the scheduled
job. Note that this definition implicitly assumes that the scheduler is work-conserving in
the sense of the definition given in prosa.classic.model.schedule.uni.basic.platform. Therefore, it
cannot be applied to models with jitter or self-suspensions. *)
Definition is_priority_inversion t :=
if sched t is Some jlp then
~~ higher_eq_priority jlp j
else false.
(* Then we compute the cumulative priority inversion incurred by
a job within some time interval t1, t2). *)
Definition cumulative_priority_inversion t1 t2 :=
\sum_(t1 ≤ t < t2) is_priority_inversion t.
(* We say that priority inversion of job j is bounded by a constant B iff cumulative
priority inversion within any busy inverval prefix is bounded by B. *)
Definition priority_inversion_of_job_is_bounded_by (B: time) :=
∀ (t1 t2: time),
busy_interval_prefix j t1 t2 →
cumulative_priority_inversion t1 t2 ≤ B.
End JobPriorityInversionBound.
(* In this section, we define a notion of the bounded priority inversion for task. *)
Section TaskPriorityInversionBound.
(* Consider an arbitrary task tsk. *)
Variable tsk: Task.
(* We say that task tsk has bounded priority inversion if all
its jobs have bounded cumulative priority inversion. *)
Definition priority_inversion_is_bounded_by (B: time) :=
∀ (j: Job),
arrives_in arr_seq j →
job_task j = tsk →
job_cost j > 0 →
priority_inversion_of_job_is_bounded_by j B.
End TaskPriorityInversionBound.
(* In this section we define the computational
version of the notion of quiet time. *)
Section DecidableQuietTime.
(* We say that t is a quiet time for j iff every higher-priority job from
the arrival sequence that arrived before t has completed by that time. *)
Definition quiet_time_dec (j : Job) (t : time) :=
all
(fun j_hp ⇒ higher_eq_priority j_hp j ==> (completed_by job_cost sched j_hp t))
(jobs_arrived_before arr_seq t).
(* We also show that the computational and propositional definitions are equivalent. *)
Lemma quiet_time_P :
∀ j t, reflect (quiet_time j t) (quiet_time_dec j t).
End DecidableQuietTime.
(* Now we prove some lemmas about busy intervals. *)
Section Lemmas.
(* Consider an arbitrary task tsk. *)
Variable tsk: Task.
(* Consider an arbitrary job j. *)
Variable j: Job.
Hypothesis H_from_arrival_sequence: arrives_in arr_seq j.
Hypothesis H_job_task: job_task j = tsk.
Hypothesis H_job_cost_positive: job_cost_positive job_cost j.
(* Recall the list of jobs that arrive in any interval. *)
Let quiet_time t1 := quiet_time j t1.
Let busy_interval_prefix t1 t2 := busy_interval_prefix j t1 t2.
Let busy_interval t1 t2 := busy_interval j t1 t2.
Let is_priority_inversion_bounded_by K := priority_inversion_of_job_is_bounded_by j K.
(* We begin by proving a basic lemma about completion of the job within its busy interval. *)
Section BasicLemma.
(* Assume that the priority relation is reflexive. *)
Hypothesis H_priority_is_reflexive: FP_is_reflexive higher_eq_priority.
(* Consider any busy interval t1, t2) of job j. *)
Variable t1 t2: time.
Hypothesis H_busy_interval: busy_interval t1 t2.
(* We prove that job j completes by the end of the busy interval. *)
Lemma job_completes_within_busy_interval:
job_completed_by j t2.
End BasicLemma.
(* In this section, we prove that during a busy interval there
always exists a pending job. *)
Section ExistsPendingJob.
(* Assume that jobs do not execute after completion. *)
Hypothesis H_completed_jobs_dont_execute:
completed_jobs_dont_execute job_cost sched.
(* Let t1, t2 be any interval where time t1 is quiet and time t2 is not quiet. *)
Variable t1 t2: time.
Hypothesis H_interval: t1 ≤ t2.
Hypothesis H_quiet: quiet_time t1.
Hypothesis H_not_quiet: ¬ quiet_time t2.
(* Then, we prove that there is a job pending at time t2
that has higher or equal priority (with respect of tsk). *)
Lemma not_quiet_implies_exists_pending_job:
∃ j_hp,
arrives_in arr_seq j_hp ∧
arrived_between job_arrival j_hp t1 t2 ∧
higher_eq_priority j_hp j ∧
¬ job_completed_by j_hp t2.
End ExistsPendingJob.
(* In this section, we prove that during a busy interval the
processor is never idle. *)
Section ProcessorAlwaysBusy.
(* Assume that the schedule is work-conserving and that jobs do
not execute before their arrival nor after completion. *)
Hypothesis H_work_conserving: work_conserving job_arrival job_cost arr_seq sched.
Hypothesis H_completed_jobs_dont_execute: completed_jobs_dont_execute job_cost sched.
Hypothesis H_jobs_must_arrive_to_execute: jobs_must_arrive_to_execute job_arrival sched.
(* Next, we assume that the priority relation is reflexive and transitive. *)
Hypothesis H_priority_is_reflexive: JLFP_is_reflexive higher_eq_priority.
Hypothesis H_priority_is_transitive: JLFP_is_transitive higher_eq_priority.
(* Consider any busy interval prefix t1, t2). *)
Variable t1 t2: time.
Hypothesis H_busy_interval_prefix: busy_interval_prefix t1 t2.
(* We prove that if the processot is idle at a time instant t, then
the next time instant t+1 will be a quiet time. *)
Lemma idle_time_implies_quiet_time_at_the_next_time_instant:
∀ t,
is_idle sched t →
quiet_time t.+1.
(* Next, we prove that at any time instant t within the busy interval there exists a job
jhp such that (1) job jhp is pending at time t and (2) job jhp has higher-or-equal
priority than task tsk. *)
Lemma pending_hp_job_exists:
∀ t,
t1 ≤ t < t2 →
∃ jhp,
arrives_in arr_seq jhp ∧
job_pending_at jhp t ∧
higher_eq_priority jhp j.
(* We prove that at any time instant t within t1, t2) the processor is not idle. *)
Lemma not_quiet_implies_not_idle:
∀ t,
t1 ≤ t < t2 →
¬ is_idle sched t.
End ProcessorAlwaysBusy.
(* In section we prove a few auxiliary lemmas about quiet time and service. *)
Section QuietTimeAndServiceOfJobs.
(* Assume that there are no duplicate job arrivals... *)
Hypothesis H_arrival_sequence_is_a_set:
arrival_sequence_is_a_set arr_seq.
(* ...and that 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.
(* We also assume that the schedule is work-conserving. *)
Hypothesis H_work_conserving: work_conserving job_arrival job_cost arr_seq sched.
(* Let t1 be a quiet time. *)
Variable t1: time.
Hypothesis H_quiet_time: quiet_time t1.
(* Assume that there is no quiet time in the interval (t1, t1 + Δ]. *)
Variable Δ: time.
Hypothesis H_no_quiet_time: ∀ t, t1 < t ≤ t1 + Δ → ¬ quiet_time t.
(* For clarity, we introduce a notion of the total service of jobs released in
time interval t_beg, t_end) during the time interval [t1, t1 + Δ). *)
Let service_received_by_hep_jobs_released_during t_beg t_end :=
service_of_higher_or_equal_priority_jobs
sched (arrivals_between t_beg t_end) higher_eq_priority j t1 (t1 + Δ).
(* We prove that jobs with higher-than-or-equal priority that
released before time instant t1 receive no service after
time instant t1. *)
Lemma hep_jobs_receive_no_service_before_quiet_time:
service_received_by_hep_jobs_released_during t1 (t1 + Δ) =
service_received_by_hep_jobs_released_during 0 (t1 + Δ).
(* Next we prove that the total service within a "non-quiet"
time interval t1, t1 + Δ) is exactly Δ. *)
Lemma no_idle_time_within_non_quiet_time_interval:
service_of_jobs sched (arrivals_between 0 (t1 + Δ)) predT t1 (t1 + Δ) = Δ.
End QuietTimeAndServiceOfJobs.
(* In this section, we show that the length of any busy interval
is bounded, as long as there is enough supply to accomodate
the workload of tasks with higher or equal priority. *)
Section BoundingBusyInterval.
(* Assume that there are no duplicate job arrivals... *)
Hypothesis H_arrival_sequence_is_a_set:
arrival_sequence_is_a_set arr_seq.
(* ...and that 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.
(* Also assume a work-conserving JLFP schedule, ... *)
Hypothesis H_work_conserving: work_conserving job_arrival job_cost arr_seq sched.
(* ...in which the priority relation is reflexive and transitive. *)
Hypothesis H_priority_is_reflexive: JLFP_is_reflexive higher_eq_priority.
Hypothesis H_priority_is_transitive: JLFP_is_transitive higher_eq_priority.
(* Next, we recall the notion of workload of all jobs released in a given interval
t1, t2) that have higher-or-equal priority w.r.t the job j being analyzed. *)
Let hp_workload t1 t2 :=
workload_of_higher_or_equal_priority_jobs
job_cost (arrivals_between t1 t2) higher_eq_priority j.
(* With regard to the jobs with higher-or-equal priority that are released
in a given interval t1, t2), we also recall the service received by these jobs in the same interval [t1, t2). *)
Let hp_service t1 t2 :=
service_of_higher_or_equal_priority_jobs
sched (arrivals_between t1 t2) higher_eq_priority j t1 t2.
(* Now we begin the proof. First, we show that the busy interval is bounded. *)
Section BoundingBusyInterval.
(* Suppose that job j is pending at time t_busy. *)
Variable t_busy: time.
Hypothesis H_j_is_pending: job_pending_at j t_busy.
(* First, we show that there must exist a busy interval prefix. *)
Section LowerBound.
(* Since job j is pending, there is a (potentially unbounded)
busy interval that starts no later than with the arrival of j. *)
Lemma exists_busy_interval_prefix:
∃ t1,
busy_interval_prefix t1 t_busy.+1 ∧
t1 ≤ job_arrival j ≤ t_busy.
End LowerBound.
(* Next we prove that, if there is a point where the requested workload
is upper-bounded by the supply, then the busy interval eventually ends. *)
Section UpperBound.
(* Consider any busy interval prefix of job j. *)
Variable t1: time.
Hypothesis H_is_busy_prefix: busy_interval_prefix t1 t_busy.+1.
(* Let priority_inversion_bound be a constant which bounds
length of a priority inversion. *)
Variable priority_inversion_bound: time.
Hypothesis H_priority_inversion_is_bounded:
is_priority_inversion_bounded_by priority_inversion_bound.
(* Next, assume that for some positive delta, the sum of requested workload
at time t1 + delta and constant priority_inversion_bound is bounded by
delta (i.e., the supply). *)
Variable delta: time.
Hypothesis H_delta_positive: delta > 0.
Hypothesis H_workload_is_bounded:
priority_inversion_bound + hp_workload t1 (t1 + delta) ≤ delta.
(* If there is a quiet time by time (t1 + delta), it trivially follows that
the busy interval is bounded.
Thus, let's consider first the harder case where there is no quiet time,
which turns out to be impossible. *)
Section CannotBeBusyForSoLong.
(* Assume that there is no quiet time in the interval (t1, t1 + delta]. *)
Hypothesis H_no_quiet_time:
∀ t, t1 < t ≤ t1 + delta → ¬ quiet_time t.
(* Since the interval is always non-quiet, the processor is always busy
with tasks of higher-or-equal priority or some lower priority job which was scheduled,
i.e., the sum of service done by jobs with actual arrival time in t1, t1 + delta) and priority inversion equals delta. *)
Lemma busy_interval_has_uninterrupted_service:
delta ≤ priority_inversion_bound + hp_service t1 (t1 + delta).
(* Moreover, the fact that the interval is not quiet also implies
that there's more workload requested than service received. *)
Lemma busy_interval_too_much_workload:
hp_workload t1 (t1 + delta) > hp_service t1 (t1 + delta).
(* Using the two lemmas above, we infer that the workload is larger than the
interval length. However, this contradicts the assumption H_workload_is_bounded. *)
Corollary busy_interval_workload_larger_than_interval:
priority_inversion_bound + hp_workload t1 (t1 + delta) > delta.
End CannotBeBusyForSoLong.
(* Since the interval cannot remain busy for so long, we prove that
the busy interval finishes at some point t2 <= t1 + delta. *)
Lemma busy_interval_is_bounded:
∃ t2,
t2 ≤ t1 + delta ∧
busy_interval t1 t2.
End UpperBound.
End BoundingBusyInterval.
(* In this section, we show that from a workload bound we can infer
the existence of a busy interval. *)
Section BusyIntervalFromWorkloadBound.
(* Let priority_inversion_bound be a constant that bounds the length of a priority inversion. *)
Variable priority_inversion_bound: time.
Hypothesis H_priority_inversion_is_bounded:
is_priority_inversion_bounded_by priority_inversion_bound.
(* Assume that for some positive delta, the sum of requested workload at
time (t1 + delta) and priority inversion is bounded by delta (i.e., the supply). *)
Variable delta: time.
Hypothesis H_delta_positive: delta > 0.
Hypothesis H_workload_is_bounded:
∀ t, priority_inversion_bound + hp_workload t (t + delta) ≤ delta.
(* Next, we assume that job j has positive cost, from which we can
infer that there is a time in which j is pending. *)
Hypothesis H_positive_cost: job_cost j > 0.
(* Therefore there must exists a busy interval t1, t2) that contains the arrival time of j. *)
Corollary exists_busy_interval:
∃ t1 t2,
t1 ≤ job_arrival j < t2 ∧
t2 ≤ t1 + delta ∧
busy_interval t1 t2.
End BusyIntervalFromWorkloadBound.
(* If we know that the workload is bounded, we can also use the
busy interval to infer a response-time bound. *)
Section ResponseTimeBoundFromBusyInterval.
(* Let priority_inversion_bound be a constant that bounds the length of a priority inversion. *)
Variable priority_inversion_bound: time.
Hypothesis H_priority_inversion_is_bounded:
is_priority_inversion_bounded_by priority_inversion_bound.
(* Assume that for some positive delta, the sum of requested workload at
time (t1 + delta) and priority inversion is bounded by delta (i.e., the supply). *)
Variable delta: time.
Hypothesis H_delta_positive: delta > 0.
Hypothesis H_workload_is_bounded:
∀ t, priority_inversion_bound + hp_workload t (t + delta) ≤ delta.
(* Then, job j must complete by (job_arrival j + delta). *)
Lemma busy_interval_bounds_response_time:
job_completed_by j (job_arrival j + delta).
End ResponseTimeBoundFromBusyInterval.
End BoundingBusyInterval.
End Lemmas.
(* In this section, we derive an alternative condition for the existence of
a busy interval. The new condition requires the total workload (instead
of the high-priority workload) generated by the task set to be bounded. *)
Section NonOverloadedProcessor.
(* The processor has no carry-in at time t iff every job (regardless of priority)
from the arrival sequence released before t has completed by that time. *)
Definition no_carry_in (t: time) :=
∀ j_o,
arrives_in arr_seq j_o →
arrived_before job_arrival j_o t →
job_completed_by j_o t.
(* The fact that there is no carry-in at time instant t
trivially implies that t is a quiet time. *)
Lemma no_carry_in_implies_quiet_time :
∀ j t,
no_carry_in t →
quiet_time j t.
(* Assume that there are no duplicate job arrivals. *)
Hypothesis H_arrival_sequence_is_a_set:
arrival_sequence_is_a_set arr_seq.
(* We also assume that the schedule is work-conserving and that jobs
do not execute before their arrival nor after completion. *)
Hypothesis H_work_conserving: work_conserving job_arrival job_cost arr_seq sched.
Hypothesis H_completed_jobs_dont_execute: completed_jobs_dont_execute job_cost sched.
Hypothesis H_jobs_must_arrive_to_execute: jobs_must_arrive_to_execute job_arrival sched.
(* Next we show that an idle time implies no carry in at this time instant. *)
Lemma idle_instant_implies_no_carry_in_at_t :
∀ t,
is_idle sched t →
no_carry_in t.
(* Moreover, an idle time implies no carry in at the next time instant. *)
Lemma idle_instant_implies_no_carry_in_at_t_pl_1 :
∀ t,
is_idle sched t →
no_carry_in t.+1.
(* Let the priority relation be reflexive. *)
Hypothesis H_priority_is_reflexive: JLFP_is_reflexive higher_eq_priority.
(* Next, we recall the notion of workload of all jobs released in a given interval t1, t2)... *)
Let total_workload t1 t2 :=
workload_of_jobs job_cost (arrivals_between t1 t2) predT.
(* ... and total service of jobs within some time interval t1, t2). *)
Let total_service t1 t2 :=
service_of_jobs sched (arrivals_between 0 t2) predT t1 t2.
(* Assume that for some positive Δ, the sum of requested workload
at time (t + Δ) is bounded by delta (i.e., the supply).
Note that this assumption bounds the total workload of
jobs released in a time interval t, t + Δ) regardless of their priorities. *)
Variable Δ: time.
Hypothesis H_delta_positive: Δ > 0.
Hypothesis H_workload_is_bounded: ∀ t, total_workload t (t + Δ) ≤ Δ.
(* Next we prove that since for any time instant t there is a point where
the total workload is upper-bounded by the supply the processor encounters
no carry-in instants at least every Δ time units. *)
Section ProcessorIsNotTooBusy.
(* We start by proving that the processor has no carry-in at
the beginning (i.e., has no carry-in at time instant 0). *)
Lemma no_carry_in_at_the_beginning :
no_carry_in 0.
(* In this section, we prove that for any time instant t there
exists another time instant t' ∈ (t, t + Δ] such that the
processor has no carry-in at time t'. *)
Section ProcessorIsNotTooBusyInduction.
(* Consider an arbitrary time instant t... *)
Variable t: time.
(* ...such that the processor has no carry-in at time t. *)
Hypothesis H_no_carry_in: no_carry_in t.
(* First, recall that the total service is bounded by the total workload. Therefore
the total service of jobs in the interval t, t + Δ) is bounded by Δ. *)
Lemma total_service_is_bounded_by_Δ :
total_service t (t + Δ) ≤ Δ.
(* Next we consider two cases:
(1) The case when the total service is strictly less than Δ,
(2) And when the total service is equal to Δ. *)
(* In the first case, we use the pigeonhole principle to conclude
that there is an idle time instant; which in turn implies existence
of a time instant with no carry-in. *)
Lemma low_total_service_implies_existence_of_time_with_no_carry_in :
total_service t (t + Δ) < Δ →
∃ δ, δ < Δ ∧ no_carry_in (t.+1 + δ).
(* In the second case, the total service within the time interval t, t + Δ) is equal to Δ. On the other hand, we know that the total workload is lower-bounded by the total service and upper-bounded by Δ. Therefore, the total workload is equal to total service this implies completion of all jobs by time [t + Δ] and hence no carry-in at time [t + Δ]. *)
Lemma completion_of_all_jobs_implies_no_carry_in :
total_service t (t + Δ) = Δ →
no_carry_in (t + Δ).
End ProcessorIsNotTooBusyInduction.
(* Finally, we show that any interval of length Δ contains a time instant with no carry-in. *)
Lemma processor_is_not_too_busy :
∀ t, ∃ δ, δ < Δ ∧ no_carry_in (t + δ).
End ProcessorIsNotTooBusy.
(* Consider an arbitrary job j with positive cost. *)
Variable j: Job.
Hypothesis H_from_arrival_sequence: arrives_in arr_seq j.
Hypothesis H_job_cost_positive: job_cost_positive job_cost j.
(* We show that there must exists a busy interval t1, t2) that contains the arrival time of j. *)
Corollary exists_busy_interval_from_total_workload_bound :
∃ t1 t2,
t1 ≤ job_arrival j < t2 ∧
t2 ≤ t1 + Δ ∧
busy_interval j t1 t2.
End NonOverloadedProcessor.
End Definitions.
End BusyIntervalJLFP.