Library prosa.analysis.facts.model.scheduled
Require Export prosa.model.schedule.scheduled.
Require Export prosa.model.processor.platform_properties.
Require Export prosa.analysis.facts.behavior.arrivals.
Require Export prosa.model.processor.platform_properties.
Require Export prosa.analysis.facts.behavior.arrivals.
Correctness of the Scheduled Job(s)
Consider any type of jobs with arrival times.
Next, consider any kind of processor state model, ...
... any valid arrival sequence, ...
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 any schedule.
Now we establish the validity conditions under which these definitions
yield the expected results: all jobs must come from the arrival sequence
and do not execute before their arrival.
NB: We do not use valid_schedule here to avoid introducing a dependency
on a readiness mode.
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.
jobs_come_from_arrival_sequence sched arr_seq.
Hypothesis H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched.
The Set of Jobs Scheduled at a Given Time
Lemma scheduled_jobs_at_iff :
∀ j t,
j \in scheduled_jobs_at arr_seq sched t = scheduled_at sched j t.
Proof.
move⇒ j t; rewrite mem_filter; apply: andb_idr ⇒ SCHED.
by apply: arrivals_before_scheduled_at; rt_eauto.
Qed.
∀ j t,
j \in scheduled_jobs_at arr_seq sched t = scheduled_at sched j t.
Proof.
move⇒ j t; rewrite mem_filter; apply: andb_idr ⇒ SCHED.
by apply: arrivals_before_scheduled_at; rt_eauto.
Qed.
Conversely, if no jobs are in the sequence, then no job is scheduled.
Lemma scheduled_jobs_at_nil :
∀ t,
nilp (scheduled_jobs_at arr_seq sched t)
↔ (∀ j, ~~ scheduled_at sched j t).
Proof.
move⇒ t; split ⇒ [/nilP NIL j|NS]
; first by apply: contraT ⇒ /negbNE
; rewrite -scheduled_jobs_at_iff NIL in_nil.
apply /nilP.
case NN: (scheduled_jobs_at _) ⇒ [//|j js].
exfalso.
have: scheduled_at sched j t; last by move: (NS j) ⇒ /negP.
rewrite -scheduled_jobs_at_iff NN.
exact: mem_head.
Qed.
∀ t,
nilp (scheduled_jobs_at arr_seq sched t)
↔ (∀ j, ~~ scheduled_at sched j t).
Proof.
move⇒ t; split ⇒ [/nilP NIL j|NS]
; first by apply: contraT ⇒ /negbNE
; rewrite -scheduled_jobs_at_iff NIL in_nil.
apply /nilP.
case NN: (scheduled_jobs_at _) ⇒ [//|j js].
exfalso.
have: scheduled_at sched j t; last by move: (NS j) ⇒ /negP.
rewrite -scheduled_jobs_at_iff NN.
exact: mem_head.
Qed.
The Job Scheduled on a Uniprocessor
Suppose we are looking at a uniprocessor.
Then clearly there is at most one job scheduled at any time t.
Corollary scheduled_jobs_at_seq1 :
∀ t,
size (scheduled_jobs_at arr_seq sched t) ≤ 1.
Proof.
move⇒ t; set sja := scheduled_jobs_at _ _ _; rewrite leqNgt.
have uniq_sja : uniq sja by apply/filter_uniq/arrivals_uniq; rt_eauto.
apply/negP ⇒ /exists_two/(_ uniq_sja)[j1 [j2 [j1j2 [j1sja j2sja]]]].
by apply/j1j2/(H_uni j1 j2 sched t); rewrite -scheduled_jobs_at_iff.
Qed.
∀ t,
size (scheduled_jobs_at arr_seq sched t) ≤ 1.
Proof.
move⇒ t; set sja := scheduled_jobs_at _ _ _; rewrite leqNgt.
have uniq_sja : uniq sja by apply/filter_uniq/arrivals_uniq; rt_eauto.
apply/negP ⇒ /exists_two/(_ uniq_sja)[j1 [j2 [j1j2 [j1sja j2sja]]]].
by apply/j1j2/(H_uni j1 j2 sched t); rewrite -scheduled_jobs_at_iff.
Qed.
Corollary scheduled_job_at_iff :
∀ j t,
scheduled_job_at arr_seq sched t = Some j ↔ scheduled_at sched j t.
Proof.
move⇒ j t; rewrite -scheduled_jobs_at_iff /scheduled_job_at.
have := scheduled_jobs_at_seq1 t.
case: scheduled_jobs_at ⇒ [//|j' [|//]] _ /=.
by rewrite in_cons in_nil orbF; split ⇒ [[]->|/eqP ->].
Qed.
∀ j t,
scheduled_job_at arr_seq sched t = Some j ↔ scheduled_at sched j t.
Proof.
move⇒ j t; rewrite -scheduled_jobs_at_iff /scheduled_job_at.
have := scheduled_jobs_at_seq1 t.
case: scheduled_jobs_at ⇒ [//|j' [|//]] _ /=.
by rewrite in_cons in_nil orbF; split ⇒ [[]->|/eqP ->].
Qed.
Corollary scheduled_job_at_none :
∀ t,
scheduled_job_at arr_seq sched t = None
↔ (∀ j, ~~ scheduled_at sched j t).
Proof.
move⇒ t. rewrite /scheduled_job_at -scheduled_jobs_at_nil.
by case: (scheduled_jobs_at arr_seq sched t).
Qed.
End Uniprocessors.
∀ t,
scheduled_job_at arr_seq sched t = None
↔ (∀ j, ~~ scheduled_at sched j t).
Proof.
move⇒ t. rewrite /scheduled_job_at -scheduled_jobs_at_nil.
by case: (scheduled_jobs_at arr_seq sched t).
Qed.
End Uniprocessors.
Case Analysis: a Scheduled Job Exists or no Job is Scheduled
Lemma scheduled_at_dec :
∀ t,
{∃ j, scheduled_at sched j t} + {∀ j, ~~ scheduled_at sched j t}.
Proof.
move⇒ t.
case SJA: (scheduled_jobs_at arr_seq sched t) ⇒ [|j js].
- by right; rewrite -scheduled_jobs_at_nil SJA.
- left; ∃ j.
rewrite -scheduled_jobs_at_iff SJA.
exact: mem_head.
Qed.
End ScheduledJobs.
∀ t,
{∃ j, scheduled_at sched j t} + {∀ j, ~~ scheduled_at sched j t}.
Proof.
move⇒ t.
case SJA: (scheduled_jobs_at arr_seq sched t) ⇒ [|j js].
- by right; rewrite -scheduled_jobs_at_nil SJA.
- left; ∃ j.
rewrite -scheduled_jobs_at_iff SJA.
exact: mem_head.
Qed.
End ScheduledJobs.