Built with Alectryon, running Coq+SerAPI v8.20.0+0.20.0. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use ⌘ instead of Ctrl.
[Loading ML file ssrmatching_plugin.cmxs (using legacy method) ... done]
[Loading ML file ssreflect_plugin.cmxs (using legacy method) ... done]
[Loading ML file ring_plugin.cmxs (using legacy method) ... done]
Serlib plugin: coq-elpi.elpi is not available: serlib support is missing.
Incremental checking for commands in this plugin will be impacted.
[Loading ML file coq-elpi.elpi ... done]
[Loading ML file zify_plugin.cmxs (using legacy method) ... done]
[Loading ML file micromega_core_plugin.cmxs (using legacy method) ... done]
[Loading ML file micromega_plugin.cmxs (using legacy method) ... done]
[Loading ML file btauto_plugin.cmxs (using legacy method) ... done]
Notation"_ + _" was already used in scope nat_scope.
[notation-overridden,parsing,default]
Notation"_ - _" was already used in scope nat_scope.
[notation-overridden,parsing,default]
Notation"_ <= _" was already used in scope nat_scope.
[notation-overridden,parsing,default]
Notation"_ < _" was already used in scope nat_scope.
[notation-overridden,parsing,default]
Notation"_ >= _" was already used in scope nat_scope.
[notation-overridden,parsing,default]
Notation"_ > _" was already used in scope nat_scope.
[notation-overridden,parsing,default]
Notation"_ <= _ <= _" was already used in scope
nat_scope. [notation-overridden,parsing,default]
Notation"_ < _ <= _" was already used in scope
nat_scope. [notation-overridden,parsing,default]
Notation"_ <= _ < _" was already used in scope
nat_scope. [notation-overridden,parsing,default]
Notation"_ < _ < _" was already used in scope
nat_scope. [notation-overridden,parsing,default]
Notation"_ * _" was already used in scope nat_scope.
[notation-overridden,parsing,default]
Require Import prosa.results.edf.rta.bounded_nps.Require Export prosa.analysis.facts.preemption.task.preemptive.Require Export prosa.analysis.facts.preemption.rtc_threshold.preemptive.Require Export prosa.analysis.facts.readiness.basic.Require Import prosa.model.task.preemption.fully_preemptive.Require Import prosa.model.priority.edf.(** * RTA for Fully Preemptive EDF *)(** In this section we prove the RTA theorem for the fully preemptive EDF model *)(** ** Setup and Assumptions *)SectionRTAforFullyPreemptiveEDFModelwithArrivalCurves.(** We assume that jobs and tasks are fully preemptive. *)#[local] Existing Instancefully_preemptive_job_model.#[local] Existing Instancefully_preemptive_task_model.#[local] Existing Instancefully_preemptive_rtc_threshold.(** Consider any type of tasks ... *)Context {Task : TaskType}.Context `{TaskCost Task}.Context `{TaskDeadline Task}.(** ... and any type of jobs associated with these tasks. *)Context {Job : JobType}.Context `{JobTask Job Task}.Context `{JobArrival Job}.Context `{JobCost Job}.(** We assume the classic (i.e., Liu & Layland) model of readiness without jitter or self-suspensions, wherein pending jobs are always ready. *)#[local] Existing Instancebasic_ready_instance.(** Consider any arrival sequence with consistent, non-duplicate arrivals. *)Variablearr_seq : arrival_sequence Job.HypothesisH_valid_arrival_sequence : valid_arrival_sequence arr_seq.(** Consider an arbitrary task set ts, ... *)Variablets : list Task.(** ... assume that all jobs come from this task set, ... *)HypothesisH_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts.(** ... and the cost of a job cannot be larger than the task cost. *)HypothesisH_valid_job_cost:
arrivals_have_valid_job_costs arr_seq.(** Let max_arrivals be a family of valid arrival curves, i.e., for any task [tsk] in ts [max_arrival tsk] is (1) an arrival bound of [tsk], and (2) it is a monotonic function that equals 0 for the empty interval delta = 0. *)Context `{MaxArrivals Task}.HypothesisH_valid_arrival_curve : valid_taskset_arrival_curve ts max_arrivals.HypothesisH_is_arrival_curve : taskset_respects_max_arrivals arr_seq ts.(** Let [tsk] be any task in ts that is to be analyzed. *)Variabletsk : Task.HypothesisH_tsk_in_ts : tsk \in ts.(** Next, consider any valid ideal uniprocessor schedule of the arrival sequence ... *)Variablesched : schedule (ideal.processor_state Job).HypothesisH_sched_valid: valid_schedule sched arr_seq.HypothesisH_jobs_come_from_arrival_sequence:
jobs_come_from_arrival_sequence sched arr_seq.(** Next, we assume that the schedule is a work-conserving schedule... *)HypothesisH_work_conserving : work_conserving arr_seq sched.(** ... and the schedule respects the scheduling policy. *)HypothesisH_respects_policy : respects_JLFP_policy_at_preemption_point arr_seq sched (EDF Job).(** ** Total Workload and Length of Busy Interval *)(** We introduce the abbreviation [rbf] for the task request bound function, which is defined as [task_cost(T) × max_arrivals(T,Δ)] for a task T. *)Letrbf := task_request_bound_function.(** Next, we introduce [task_rbf] as an abbreviation for the task request bound function of task [tsk]. *)Lettask_rbf := rbf tsk.(** Using the sum of individual request bound functions, we define the request bound function of all tasks (total request bound function). *)Lettotal_rbf := total_request_bound_function ts.(** If jobs are fully preemptive, lower priority jobs do not cause priority inversion. Hence, the blocking bound is always 0 for any [A]. *)Letblocking_bound (A : duration) := 0.(** Let L be any positive fixed point of the busy interval recurrence. *)VariableL : duration.HypothesisH_L_positive : L > 0.HypothesisH_fixed_point : L = total_rbf L.(** ** Response-Time Bound *)(** To reduce the time complexity of the analysis, recall the notion of search space. *)Letis_in_search_space := bounded_nps.is_in_search_space ts tsk L.(** Consider any value [R], and assume that for any given arrival offset [A] in the search space, there is a solution of the response-time bound recurrence which is bounded by [R]. *)VariableR : duration.HypothesisH_R_is_maximum:
forall (A : duration),
is_in_search_space A ->
exists (F : duration),
A + F >= task_rbf (A + ε) + bound_on_athep_workload ts tsk A (A + F) /\
R >= F.(** Now, we can leverage the results for the abstract model with bounded non-preemptive segments to establish a response-time bound for the more concrete model of fully preemptive scheduling. *)Letresponse_time_bounded_by := task_response_time_bound arr_seq sched.
forallA : duration,
bounded_nps.is_in_search_space ts tsk L A ->
existsF : duration,
edf.blocking_bound ts tsk A +
(task_request_bound_function tsk (A + 1) -
(task_cost tsk - task_rtct tsk)) +
bound_on_athep_workload ts tsk A (A + F) <= A + F /\
F + (task_cost tsk - task_rtct tsk) <= R
existsF : duration,
edf.blocking_bound ts tsk A +
(task_request_bound_function tsk (A + 1) -
(task_cost tsk - task_rtct tsk)) +
bound_on_athep_workload ts tsk A (A + F) <= A + F /\
F + (task_cost tsk - task_rtct tsk) <= R
Task: TaskType H: TaskCost Task H0: TaskDeadline Task Job: JobType H1: JobTask Job Task H2: JobArrival Job H3: JobCost Job arr_seq: arrival_sequence Job H_valid_arrival_sequence: valid_arrival_sequence
arr_seq ts: seq Task H_all_jobs_from_taskset: all_jobs_from_taskset
arr_seq ts H_valid_job_cost: arrivals_have_valid_job_costs
arr_seq H4: MaxArrivals Task H_valid_arrival_curve: valid_taskset_arrival_curve ts
max_arrivals H_is_arrival_curve: taskset_respects_max_arrivals
arr_seq ts tsk: Task H_tsk_in_ts: tsk \in ts sched: schedule (ideal.processor_state Job) H_sched_valid: valid_schedule sched arr_seq H_jobs_come_from_arrival_sequence: jobs_come_from_arrival_sequence
sched arr_seq H_work_conserving: work_conserving arr_seq sched H_respects_policy: respects_JLFP_policy_at_preemption_point
arr_seq sched
(EDF Job) rbf:= task_request_bound_function: Task -> duration -> nat task_rbf:= rbf tsk: duration -> nat total_rbf:= total_request_bound_function ts: duration -> nat blocking_bound:= fun=> 0: duration -> nat L: duration H_L_positive: 0 < L H_fixed_point: L = total_rbf L is_in_search_space:= bounded_nps.is_in_search_space ts
tsk L: duration -> bool R, A: duration H_R_is_maximum: existsF : duration,
task_rbf (A + 1) +
bound_on_athep_workload ts tsk A
(A + F) <=
A + F /\
F <= R response_time_bounded_by:= task_response_time_bound
arr_seq sched: Task -> duration -> Prop LT: A < L CHANGE: task_rbf_changes_at tsk A
|| bound_on_total_hep_workload_changes_at ts
tsk A BLOCK: forallA' : duration,
edf.blocking_bound ts tsk A' =
blocking_bound A' F: duration FIX: task_rbf (A + 1) +
bound_on_athep_workload ts tsk A (A + F) <=
A + F BOUND: F <= R
existsF : duration,
edf.blocking_bound ts tsk A +
(task_request_bound_function tsk (A + 1) -
(task_cost tsk - task_rtct tsk)) +
bound_on_athep_workload ts tsk A (A + F) <= A + F /\
F + (task_cost tsk - task_rtct tsk) <= R
Task: TaskType H: TaskCost Task H0: TaskDeadline Task Job: JobType H1: JobTask Job Task H2: JobArrival Job H3: JobCost Job arr_seq: arrival_sequence Job H_valid_arrival_sequence: valid_arrival_sequence
arr_seq ts: seq Task H_all_jobs_from_taskset: all_jobs_from_taskset
arr_seq ts H_valid_job_cost: arrivals_have_valid_job_costs
arr_seq H4: MaxArrivals Task H_valid_arrival_curve: valid_taskset_arrival_curve ts
max_arrivals H_is_arrival_curve: taskset_respects_max_arrivals
arr_seq ts tsk: Task H_tsk_in_ts: tsk \in ts sched: schedule (ideal.processor_state Job) H_sched_valid: valid_schedule sched arr_seq H_jobs_come_from_arrival_sequence: jobs_come_from_arrival_sequence
sched arr_seq H_work_conserving: work_conserving arr_seq sched H_respects_policy: respects_JLFP_policy_at_preemption_point
arr_seq sched
(EDF Job) rbf:= task_request_bound_function: Task -> duration -> nat task_rbf:= rbf tsk: duration -> nat total_rbf:= total_request_bound_function ts: duration -> nat blocking_bound:= fun=> 0: duration -> nat L: duration H_L_positive: 0 < L H_fixed_point: L = total_rbf L is_in_search_space:= bounded_nps.is_in_search_space ts
tsk L: duration -> bool R, A: duration H_R_is_maximum: existsF : duration,
task_rbf (A + 1) +
bound_on_athep_workload ts tsk A
(A + F) <=
A + F /\
F <= R response_time_bounded_by:= task_response_time_bound
arr_seq sched: Task -> duration -> Prop LT: A < L CHANGE: task_rbf_changes_at tsk A
|| bound_on_total_hep_workload_changes_at ts
tsk A BLOCK: forallA' : duration,
edf.blocking_bound ts tsk A' =
blocking_bound A' F: duration FIX: task_rbf (A + 1) +
bound_on_athep_workload ts tsk A (A + F) <=
A + F BOUND: F <= R
edf.blocking_bound ts tsk A +
(task_request_bound_function tsk (A + 1) -
(task_cost tsk - task_rtct tsk)) +
bound_on_athep_workload ts tsk A (A + F) <= A + F
Task: TaskType H: TaskCost Task H0: TaskDeadline Task Job: JobType H1: JobTask Job Task H2: JobArrival Job H3: JobCost Job arr_seq: arrival_sequence Job H_valid_arrival_sequence: valid_arrival_sequence
arr_seq ts: seq Task H_all_jobs_from_taskset: all_jobs_from_taskset
arr_seq ts H_valid_job_cost: arrivals_have_valid_job_costs
arr_seq H4: MaxArrivals Task H_valid_arrival_curve: valid_taskset_arrival_curve ts
max_arrivals H_is_arrival_curve: taskset_respects_max_arrivals
arr_seq ts tsk: Task H_tsk_in_ts: tsk \in ts sched: schedule (ideal.processor_state Job) H_sched_valid: valid_schedule sched arr_seq H_jobs_come_from_arrival_sequence: jobs_come_from_arrival_sequence
sched arr_seq H_work_conserving: work_conserving arr_seq sched H_respects_policy: respects_JLFP_policy_at_preemption_point
arr_seq sched
(EDF Job) rbf:= task_request_bound_function: Task -> duration -> nat task_rbf:= rbf tsk: duration -> nat total_rbf:= total_request_bound_function ts: duration -> nat blocking_bound:= fun=> 0: duration -> nat L: duration H_L_positive: 0 < L H_fixed_point: L = total_rbf L is_in_search_space:= bounded_nps.is_in_search_space ts
tsk L: duration -> bool R, A: duration H_R_is_maximum: existsF : duration,
task_rbf (A + 1) +
bound_on_athep_workload ts tsk A
(A + F) <=
A + F /\
F <= R response_time_bounded_by:= task_response_time_bound
arr_seq sched: Task -> duration -> Prop LT: A < L CHANGE: task_rbf_changes_at tsk A
|| bound_on_total_hep_workload_changes_at ts
tsk A BLOCK: forallA' : duration,
edf.blocking_bound ts tsk A' =
blocking_bound A' F: duration FIX: task_rbf (A + 1) +
bound_on_athep_workload ts tsk A (A + F) <=
A + F BOUND: F <= R
F + (task_cost tsk - task_rtct tsk) <= R
Task: TaskType H: TaskCost Task H0: TaskDeadline Task Job: JobType H1: JobTask Job Task H2: JobArrival Job H3: JobCost Job arr_seq: arrival_sequence Job H_valid_arrival_sequence: valid_arrival_sequence
arr_seq ts: seq Task H_all_jobs_from_taskset: all_jobs_from_taskset
arr_seq ts H_valid_job_cost: arrivals_have_valid_job_costs
arr_seq H4: MaxArrivals Task H_valid_arrival_curve: valid_taskset_arrival_curve ts
max_arrivals H_is_arrival_curve: taskset_respects_max_arrivals
arr_seq ts tsk: Task H_tsk_in_ts: tsk \in ts sched: schedule (ideal.processor_state Job) H_sched_valid: valid_schedule sched arr_seq H_jobs_come_from_arrival_sequence: jobs_come_from_arrival_sequence
sched arr_seq H_work_conserving: work_conserving arr_seq sched H_respects_policy: respects_JLFP_policy_at_preemption_point
arr_seq sched
(EDF Job) rbf:= task_request_bound_function: Task -> duration -> nat task_rbf:= rbf tsk: duration -> nat total_rbf:= total_request_bound_function ts: duration -> nat blocking_bound:= fun=> 0: duration -> nat L: duration H_L_positive: 0 < L H_fixed_point: L = total_rbf L is_in_search_space:= bounded_nps.is_in_search_space ts
tsk L: duration -> bool R, A: duration H_R_is_maximum: existsF : duration,
task_rbf (A + 1) +
bound_on_athep_workload ts tsk A
(A + F) <=
A + F /\
F <= R response_time_bounded_by:= task_response_time_bound
arr_seq sched: Task -> duration -> Prop LT: A < L CHANGE: task_rbf_changes_at tsk A
|| bound_on_total_hep_workload_changes_at ts
tsk A BLOCK: forallA' : duration,
edf.blocking_bound ts tsk A' =
blocking_bound A' F: duration FIX: task_rbf (A + 1) +
bound_on_athep_workload ts tsk A (A + F) <=
A + F BOUND: F <= R
edf.blocking_bound ts tsk A +
(task_request_bound_function tsk (A + 1) -
(task_cost tsk - task_rtct tsk)) +
bound_on_athep_workload ts tsk A (A + F) <= A + F
byrewrite BLOCK add0n subnn subn0.
Task: TaskType H: TaskCost Task H0: TaskDeadline Task Job: JobType H1: JobTask Job Task H2: JobArrival Job H3: JobCost Job arr_seq: arrival_sequence Job H_valid_arrival_sequence: valid_arrival_sequence
arr_seq ts: seq Task H_all_jobs_from_taskset: all_jobs_from_taskset
arr_seq ts H_valid_job_cost: arrivals_have_valid_job_costs
arr_seq H4: MaxArrivals Task H_valid_arrival_curve: valid_taskset_arrival_curve ts
max_arrivals H_is_arrival_curve: taskset_respects_max_arrivals
arr_seq ts tsk: Task H_tsk_in_ts: tsk \in ts sched: schedule (ideal.processor_state Job) H_sched_valid: valid_schedule sched arr_seq H_jobs_come_from_arrival_sequence: jobs_come_from_arrival_sequence
sched arr_seq H_work_conserving: work_conserving arr_seq sched H_respects_policy: respects_JLFP_policy_at_preemption_point
arr_seq sched
(EDF Job) rbf:= task_request_bound_function: Task -> duration -> nat task_rbf:= rbf tsk: duration -> nat total_rbf:= total_request_bound_function ts: duration -> nat blocking_bound:= fun=> 0: duration -> nat L: duration H_L_positive: 0 < L H_fixed_point: L = total_rbf L is_in_search_space:= bounded_nps.is_in_search_space ts
tsk L: duration -> bool R, A: duration H_R_is_maximum: existsF : duration,
task_rbf (A + 1) +
bound_on_athep_workload ts tsk A
(A + F) <=
A + F /\
F <= R response_time_bounded_by:= task_response_time_bound
arr_seq sched: Task -> duration -> Prop LT: A < L CHANGE: task_rbf_changes_at tsk A
|| bound_on_total_hep_workload_changes_at ts
tsk A BLOCK: forallA' : duration,
edf.blocking_bound ts tsk A' =
blocking_bound A' F: duration FIX: task_rbf (A + 1) +
bound_on_athep_workload ts tsk A (A + F) <=
A + F BOUND: F <= R