Library rt.restructuring.analysis.fixed_priority.rta.nonpr_reg.concrete_models.limited
(* ----------------------------------[ coqtop ]---------------------------------
Welcome to Coq 8.10.1 (October 2019)
----------------------------------------------------------------------------- *)
From rt.util Require Import all.
From rt.restructuring.behavior Require Import all.
From rt.restructuring.analysis.basic_facts Require Import all.
From rt.restructuring.model Require Import job task workload processor.ideal readiness.basic.
From rt.restructuring.model.arrival Require Import arrival_curves.
From rt.restructuring.model Require Import preemption.limited.
From rt.restructuring.model.schedule Require Import
work_conserving priority_based.priorities priority_based.preemption_aware.
From rt.restructuring.analysis.arrival Require Import workload_bound rbf.
From rt.restructuring.analysis.fixed_priority.rta Require Import nonpr_reg.response_time_bound.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq path fintype bigop.
RTA for FP-schedulers with Fixed Premption Points
In this module we prove the RTA theorem for FP-schedulers with fixed preemption points.
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, non-duplicate arrivals. 
  Variable arr_seq : arrival_sequence Job.
Hypothesis H_arrival_times_are_consistent : consistent_arrival_times arr_seq.
Hypothesis H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq.
Hypothesis H_arrival_times_are_consistent : consistent_arrival_times arr_seq.
Hypothesis H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq.
Consider an arbitrary task set ts, ... 
... assume that all jobs come from the task set, ... 
... and the cost of a job cannot be larger than the task cost. 
First, we assume we have the model with fixed preemption points.
      I.e., each task is divided into a number of nonpreemptive segments 
      by inserting staticaly predefined preemption points. 
  Context `{JobPreemptionPoints Job}
`{TaskPreemptionPoints Task}.
Hypothesis H_valid_model_with_fixed_preemption_points:
valid_fixed_preemption_points_model arr_seq ts.
`{TaskPreemptionPoints Task}.
Hypothesis H_valid_model_with_fixed_preemption_points:
valid_fixed_preemption_points_model arr_seq ts.
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}.
Hypothesis H_valid_arrival_curve : valid_taskset_arrival_curve ts max_arrivals.
Hypothesis H_is_arrival_curve : taskset_respects_max_arrivals arr_seq ts.
Hypothesis H_valid_arrival_curve : valid_taskset_arrival_curve ts max_arrivals.
Hypothesis H_is_arrival_curve : taskset_respects_max_arrivals arr_seq ts.
Let tsk be any task in ts that is to be analyzed. 
Next, consider any ideal uniprocessor schedule  with limited preemptionsof 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_valid_schedule_with_limited_preemptions:
valid_schedule_with_limited_preemptions arr_seq sched.
Hypothesis H_jobs_come_from_arrival_sequence:
jobs_come_from_arrival_sequence sched arr_seq.
Hypothesis H_valid_schedule_with_limited_preemptions:
valid_schedule_with_limited_preemptions arr_seq sched.
... 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.
Consider an FP policy that indicates a higher-or-equal priority relation,
     and assume that the relation is reflexive and transitive. 
  Variable higher_eq_priority : FP_policy Task.
Hypothesis H_priority_is_reflexive : reflexive_priorities.
Hypothesis H_priority_is_transitive : transitive_priorities.
Hypothesis H_priority_is_reflexive : reflexive_priorities.
Hypothesis H_priority_is_transitive : transitive_priorities.
Assume we have sequential tasks, i.e, jobs from the 
     same task execute in the order of their arrival. 
Next, we assume that the schedule is a work-conserving schedule... 
... and the schedule respects the policy defined by thejob_preemptable 
     function (i.e., jobs have bounded nonpreemptive segments). 
Let's define some local names for clarity. 
  Let task_rbf := task_request_bound_function tsk.
Let total_hep_rbf := total_hep_request_bound_function_FP _ ts tsk.
Let total_ohep_rbf := total_ohep_request_bound_function_FP _ ts tsk.
Let response_time_bounded_by := task_response_time_bound arr_seq sched.
Let total_hep_rbf := total_hep_request_bound_function_FP _ ts tsk.
Let total_ohep_rbf := total_ohep_request_bound_function_FP _ ts tsk.
Let response_time_bounded_by := task_response_time_bound arr_seq sched.
Next, we define a bound for the priority inversion caused by tasks of lower priority. 
  Let blocking_bound :=
\max_(tsk_other <- ts | ~~ higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε).
\max_(tsk_other <- ts | ~~ higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε).
Let L be any positive fixed point of the busy interval recurrence, determined by 
     the sum of blocking and higher-or-equal-priority workload. 
  Variable L : duration.
Hypothesis H_L_positive : L > 0.
Hypothesis H_fixed_point : L = blocking_bound + total_hep_rbf L.
Hypothesis H_L_positive : L > 0.
Hypothesis H_fixed_point : L = blocking_bound + total_hep_rbf L.
To reduce the time complexity of the analysis, recall the notion of search space. 
Next, consider any value R, and assume that for any given arrival A from search space
      there is a solution of the response-time bound recurrence which is bounded by R. 
  Variable R: nat.
Hypothesis H_R_is_maximum:
∀ (A : duration),
is_in_search_space A →
∃ (F : duration),
A + F = blocking_bound
+ (task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε))
+ total_ohep_rbf (A + F) ∧
F + (task_last_nonpr_segment tsk - ε) ≤ R.
Hypothesis H_R_is_maximum:
∀ (A : duration),
is_in_search_space A →
∃ (F : duration),
A + F = blocking_bound
+ (task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε))
+ total_ohep_rbf (A + F) ∧
F + (task_last_nonpr_segment tsk - ε) ≤ R.
Now, we can reuse the results for the abstract model with bounded nonpreemptive segments
     to establish a response-time bound for the more concrete model of fixed preemption points. 
  Theorem uniprocessor_response_time_bound_fp_with_fixed_preemption_points:
response_time_bounded_by tsk R.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 286)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
============================
response_time_bounded_by tsk R
----------------------------------------------------------------------------- *)
Proof.
move: (H_valid_model_with_fixed_preemption_points) ⇒ [MLP [BEG [END [INCR [HYP1 [HYP2 HYP3]]]]]].
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 348)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
============================
response_time_bounded_by tsk R
----------------------------------------------------------------------------- *)
move: (MLP) ⇒ [BEGj [ENDj _]].
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 371)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
============================
response_time_bounded_by tsk R
----------------------------------------------------------------------------- *)
edestruct (posnP (task_cost tsk)) as [ZERO|POSt].
(* ----------------------------------[ coqtop ]---------------------------------
2 subgoals (ID 396)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
ZERO : task_cost tsk = 0
============================
response_time_bounded_by tsk R
subgoal 2 (ID 397) is:
response_time_bounded_by tsk R
----------------------------------------------------------------------------- *)
{
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 396)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
ZERO : task_cost tsk = 0
============================
response_time_bounded_by tsk R
----------------------------------------------------------------------------- *)
intros j ARR TSK.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 401)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
ZERO : task_cost tsk = 0
j : Job
ARR : arrives_in arr_seq j
TSK : job_task j = tsk
============================
job_response_time_bound sched j R
----------------------------------------------------------------------------- *)
move: (H_job_cost_le_task_cost _ ARR) ⇒ POSt.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 405)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
ZERO : task_cost tsk = 0
j : Job
ARR : arrives_in arr_seq j
TSK : job_task j = tsk
POSt : job_cost_le_task_cost j
============================
job_response_time_bound sched j R
----------------------------------------------------------------------------- *)
move: POSt; rewrite /job_cost_le_task_cost TSK ZERO leqn0; move ⇒ /eqP Z.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 459)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
ZERO : task_cost tsk = 0
j : Job
ARR : arrives_in arr_seq j
TSK : job_task j = tsk
Z : job_cost j = 0
============================
job_response_time_bound sched j R
----------------------------------------------------------------------------- *)
by rewrite /job_response_time_bound /completed_by Z.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 397)
subgoal 1 (ID 397) is:
response_time_bounded_by tsk R
----------------------------------------------------------------------------- *)
}
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 397)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
============================
response_time_bounded_by tsk R
----------------------------------------------------------------------------- *)
eapply uniprocessor_response_time_bound_fp_with_bounded_nonpreemptive_segments
with (L0 := L).
(* ----------------------------------[ coqtop ]---------------------------------
21 focused subgoals
(shelved: 1) (ID 492)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
============================
consistent_arrival_times arr_seq
subgoal 2 (ID 493) is:
arrival_sequence_uniq arr_seq
subgoal 3 (ID 494) is:
jobs_come_from_arrival_sequence sched arr_seq
subgoal 4 (ID 495) is:
jobs_must_arrive_to_execute sched
subgoal 5 (ID 496) is:
completed_jobs_dont_execute sched
subgoal 6 (ID 497) is:
valid_model_with_bounded_nonpreemptive_segments arr_seq sched
subgoal 7 (ID 498) is:
reflexive_priorities
subgoal 8 (ID 499) is:
transitive_priorities
subgoal 9 (ID 500) is:
sequential_tasks sched
subgoal 10 (ID 501) is:
work_conserving arr_seq sched
subgoal 11 (ID 502) is:
respects_policy_at_preemption_point arr_seq sched
subgoal 12 (ID 503) is:
all_jobs_from_taskset arr_seq ?ts0
subgoal 13 (ID 504) is:
cost_of_jobs_from_arrival_sequence_le_task_cost arr_seq
subgoal 14 (ID 505) is:
valid_taskset_arrival_curve ?ts0 max_arrivals
subgoal 15 (ID 506) is:
taskset_respects_max_arrivals arr_seq ?ts0
subgoal 16 (ID 507) is:
tsk \in ?ts0
subgoal 17 (ID 508) is:
valid_preemption_model arr_seq sched
subgoal 18 (ID 509) is:
valid_rtct.valid_task_run_to_completion_threshold arr_seq tsk
subgoal 19 (ID 510) is:
0 < L
subgoal 20 (ID 511) is:
L =
response_time_bound.blocking_bound higher_eq_priority ?ts0 tsk +
total_hep_request_bound_function_FP higher_eq_priority ?ts0 tsk L
subgoal 21 (ID 512) is:
forall A : duration,
(A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε)) ->
exists F : duration,
A + F =
response_time_bound.blocking_bound higher_eq_priority ?ts0 tsk +
(task_request_bound_function tsk (A + ε) -
(task_cost tsk - task_run_to_completion_threshold tsk)) +
total_ohep_request_bound_function_FP higher_eq_priority ?ts0 tsk (A + F) /\
F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
----------------------------------------------------------------------------- *)
all: eauto 2 with basic_facts.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 512)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
============================
forall A : duration,
(A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε)) ->
exists F : duration,
A + F =
response_time_bound.blocking_bound higher_eq_priority ts tsk +
(task_request_bound_function tsk (A + ε) -
(task_cost tsk - task_run_to_completion_threshold tsk)) +
total_ohep_request_bound_function_FP higher_eq_priority ts tsk (A + F) /\
F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
----------------------------------------------------------------------------- *)
intros A SP.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 550)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
============================
exists F : duration,
A + F =
response_time_bound.blocking_bound higher_eq_priority ts tsk +
(task_request_bound_function tsk (A + ε) -
(task_cost tsk - task_run_to_completion_threshold tsk)) +
total_ohep_request_bound_function_FP higher_eq_priority ts tsk (A + F) /\
F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
----------------------------------------------------------------------------- *)
destruct (H_R_is_maximum _ SP) as[FF [EQ1 EQ2]].
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 562)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
exists F : duration,
A + F =
response_time_bound.blocking_bound higher_eq_priority ts tsk +
(task_request_bound_function tsk (A + ε) -
(task_cost tsk - task_run_to_completion_threshold tsk)) +
total_ohep_request_bound_function_FP higher_eq_priority ts tsk (A + F) /\
F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
----------------------------------------------------------------------------- *)
∃ FF; rewrite subKn; first by done.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 570)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
task_last_nonpr_segment tsk - ε <= task_cost tsk
----------------------------------------------------------------------------- *)
rewrite /task_last_nonpr_segment -(leq_add2r 1) subn1 !addn1 prednK; last first.
(* ----------------------------------[ coqtop ]---------------------------------
2 subgoals (ID 622)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
0 < last0 (distances (task_preemption_points tsk))
subgoal 2 (ID 621) is:
last0 (distances (task_preemption_points tsk)) <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
- rewrite /last0 -nth_last.
(* ----------------------------------[ coqtop ]---------------------------------
2 subgoals (ID 628)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
0 <
nth 0 (distances (task_preemption_points tsk))
(predn (size (distances (task_preemption_points tsk))))
subgoal 2 (ID 621) is:
last0 (distances (task_preemption_points tsk)) <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
apply HYP3; try by done.
(* ----------------------------------[ coqtop ]---------------------------------
2 subgoals (ID 630)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
predn (size (distances (task_preemption_points tsk))) <
size (distances (task_preemption_points tsk))
subgoal 2 (ID 621) is:
last0 (distances (task_preemption_points tsk)) <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
rewrite -(ltn_add2r 1) !addn1 prednK //.
(* ----------------------------------[ coqtop ]---------------------------------
2 subgoals (ID 668)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
0 < size (distances (task_preemption_points tsk))
subgoal 2 (ID 621) is:
last0 (distances (task_preemption_points tsk)) <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
move: (number_of_preemption_points_in_task_at_least_two
_ _ H_valid_model_with_fixed_preemption_points _ H_tsk_in_ts POSt) ⇒ Fact2.
(* ----------------------------------[ coqtop ]---------------------------------
2 subgoals (ID 705)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
Fact2 : 1 < size (task_preemption_points tsk)
============================
0 < size (distances (task_preemption_points tsk))
subgoal 2 (ID 621) is:
last0 (distances (task_preemption_points tsk)) <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
move: (Fact2) ⇒ Fact3.
(* ----------------------------------[ coqtop ]---------------------------------
2 subgoals (ID 707)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
Fact2, Fact3 : 1 < size (task_preemption_points tsk)
============================
0 < size (distances (task_preemption_points tsk))
subgoal 2 (ID 621) is:
last0 (distances (task_preemption_points tsk)) <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
by rewrite size_of_seq_of_distances // addn1 ltnS // in Fact2.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 621)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
last0 (distances (task_preemption_points tsk)) <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
- apply leq_trans with (task_max_nonpreemptive_segment tsk).
(* ----------------------------------[ coqtop ]---------------------------------
2 subgoals (ID 752)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
last0 (distances (task_preemption_points tsk)) <=
task_max_nonpreemptive_segment tsk
subgoal 2 (ID 753) is:
task_max_nonpreemptive_segment tsk <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
+ by apply last_of_seq_le_max_of_seq.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 753)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
task_max_nonpreemptive_segment tsk <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
+ rewrite -END; last by done.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 759)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
task_max_nonpreemptive_segment tsk <=
succn (last0 (task_preemption_points tsk))
----------------------------------------------------------------------------- *)
apply ltnW; rewrite ltnS; try done.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 765)
  
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
task_max_nonpreemptive_segment tsk <= last0 (task_preemption_points tsk)
----------------------------------------------------------------------------- *)
by apply max_distance_in_seq_le_last_element_of_seq; eauto 2.
(* ----------------------------------[ coqtop ]---------------------------------
No more subgoals.
----------------------------------------------------------------------------- *)
Qed.
End RTAforFixedPreemptionPointsModelwithArrivalCurves.
response_time_bounded_by tsk R.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 286)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
============================
response_time_bounded_by tsk R
----------------------------------------------------------------------------- *)
Proof.
move: (H_valid_model_with_fixed_preemption_points) ⇒ [MLP [BEG [END [INCR [HYP1 [HYP2 HYP3]]]]]].
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 348)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
============================
response_time_bounded_by tsk R
----------------------------------------------------------------------------- *)
move: (MLP) ⇒ [BEGj [ENDj _]].
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 371)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
============================
response_time_bounded_by tsk R
----------------------------------------------------------------------------- *)
edestruct (posnP (task_cost tsk)) as [ZERO|POSt].
(* ----------------------------------[ coqtop ]---------------------------------
2 subgoals (ID 396)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
ZERO : task_cost tsk = 0
============================
response_time_bounded_by tsk R
subgoal 2 (ID 397) is:
response_time_bounded_by tsk R
----------------------------------------------------------------------------- *)
{
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 396)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
ZERO : task_cost tsk = 0
============================
response_time_bounded_by tsk R
----------------------------------------------------------------------------- *)
intros j ARR TSK.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 401)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
ZERO : task_cost tsk = 0
j : Job
ARR : arrives_in arr_seq j
TSK : job_task j = tsk
============================
job_response_time_bound sched j R
----------------------------------------------------------------------------- *)
move: (H_job_cost_le_task_cost _ ARR) ⇒ POSt.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 405)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
ZERO : task_cost tsk = 0
j : Job
ARR : arrives_in arr_seq j
TSK : job_task j = tsk
POSt : job_cost_le_task_cost j
============================
job_response_time_bound sched j R
----------------------------------------------------------------------------- *)
move: POSt; rewrite /job_cost_le_task_cost TSK ZERO leqn0; move ⇒ /eqP Z.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 459)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
ZERO : task_cost tsk = 0
j : Job
ARR : arrives_in arr_seq j
TSK : job_task j = tsk
Z : job_cost j = 0
============================
job_response_time_bound sched j R
----------------------------------------------------------------------------- *)
by rewrite /job_response_time_bound /completed_by Z.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 397)
subgoal 1 (ID 397) is:
response_time_bounded_by tsk R
----------------------------------------------------------------------------- *)
}
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 397)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
============================
response_time_bounded_by tsk R
----------------------------------------------------------------------------- *)
eapply uniprocessor_response_time_bound_fp_with_bounded_nonpreemptive_segments
with (L0 := L).
(* ----------------------------------[ coqtop ]---------------------------------
21 focused subgoals
(shelved: 1) (ID 492)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
============================
consistent_arrival_times arr_seq
subgoal 2 (ID 493) is:
arrival_sequence_uniq arr_seq
subgoal 3 (ID 494) is:
jobs_come_from_arrival_sequence sched arr_seq
subgoal 4 (ID 495) is:
jobs_must_arrive_to_execute sched
subgoal 5 (ID 496) is:
completed_jobs_dont_execute sched
subgoal 6 (ID 497) is:
valid_model_with_bounded_nonpreemptive_segments arr_seq sched
subgoal 7 (ID 498) is:
reflexive_priorities
subgoal 8 (ID 499) is:
transitive_priorities
subgoal 9 (ID 500) is:
sequential_tasks sched
subgoal 10 (ID 501) is:
work_conserving arr_seq sched
subgoal 11 (ID 502) is:
respects_policy_at_preemption_point arr_seq sched
subgoal 12 (ID 503) is:
all_jobs_from_taskset arr_seq ?ts0
subgoal 13 (ID 504) is:
cost_of_jobs_from_arrival_sequence_le_task_cost arr_seq
subgoal 14 (ID 505) is:
valid_taskset_arrival_curve ?ts0 max_arrivals
subgoal 15 (ID 506) is:
taskset_respects_max_arrivals arr_seq ?ts0
subgoal 16 (ID 507) is:
tsk \in ?ts0
subgoal 17 (ID 508) is:
valid_preemption_model arr_seq sched
subgoal 18 (ID 509) is:
valid_rtct.valid_task_run_to_completion_threshold arr_seq tsk
subgoal 19 (ID 510) is:
0 < L
subgoal 20 (ID 511) is:
L =
response_time_bound.blocking_bound higher_eq_priority ?ts0 tsk +
total_hep_request_bound_function_FP higher_eq_priority ?ts0 tsk L
subgoal 21 (ID 512) is:
forall A : duration,
(A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε)) ->
exists F : duration,
A + F =
response_time_bound.blocking_bound higher_eq_priority ?ts0 tsk +
(task_request_bound_function tsk (A + ε) -
(task_cost tsk - task_run_to_completion_threshold tsk)) +
total_ohep_request_bound_function_FP higher_eq_priority ?ts0 tsk (A + F) /\
F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
----------------------------------------------------------------------------- *)
all: eauto 2 with basic_facts.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 512)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
============================
forall A : duration,
(A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε)) ->
exists F : duration,
A + F =
response_time_bound.blocking_bound higher_eq_priority ts tsk +
(task_request_bound_function tsk (A + ε) -
(task_cost tsk - task_run_to_completion_threshold tsk)) +
total_ohep_request_bound_function_FP higher_eq_priority ts tsk (A + F) /\
F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
----------------------------------------------------------------------------- *)
intros A SP.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 550)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
============================
exists F : duration,
A + F =
response_time_bound.blocking_bound higher_eq_priority ts tsk +
(task_request_bound_function tsk (A + ε) -
(task_cost tsk - task_run_to_completion_threshold tsk)) +
total_ohep_request_bound_function_FP higher_eq_priority ts tsk (A + F) /\
F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
----------------------------------------------------------------------------- *)
destruct (H_R_is_maximum _ SP) as[FF [EQ1 EQ2]].
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 562)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
exists F : duration,
A + F =
response_time_bound.blocking_bound higher_eq_priority ts tsk +
(task_request_bound_function tsk (A + ε) -
(task_cost tsk - task_run_to_completion_threshold tsk)) +
total_ohep_request_bound_function_FP higher_eq_priority ts tsk (A + F) /\
F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
----------------------------------------------------------------------------- *)
∃ FF; rewrite subKn; first by done.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 570)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
task_last_nonpr_segment tsk - ε <= task_cost tsk
----------------------------------------------------------------------------- *)
rewrite /task_last_nonpr_segment -(leq_add2r 1) subn1 !addn1 prednK; last first.
(* ----------------------------------[ coqtop ]---------------------------------
2 subgoals (ID 622)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
0 < last0 (distances (task_preemption_points tsk))
subgoal 2 (ID 621) is:
last0 (distances (task_preemption_points tsk)) <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
- rewrite /last0 -nth_last.
(* ----------------------------------[ coqtop ]---------------------------------
2 subgoals (ID 628)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
0 <
nth 0 (distances (task_preemption_points tsk))
(predn (size (distances (task_preemption_points tsk))))
subgoal 2 (ID 621) is:
last0 (distances (task_preemption_points tsk)) <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
apply HYP3; try by done.
(* ----------------------------------[ coqtop ]---------------------------------
2 subgoals (ID 630)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
predn (size (distances (task_preemption_points tsk))) <
size (distances (task_preemption_points tsk))
subgoal 2 (ID 621) is:
last0 (distances (task_preemption_points tsk)) <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
rewrite -(ltn_add2r 1) !addn1 prednK //.
(* ----------------------------------[ coqtop ]---------------------------------
2 subgoals (ID 668)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
0 < size (distances (task_preemption_points tsk))
subgoal 2 (ID 621) is:
last0 (distances (task_preemption_points tsk)) <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
move: (number_of_preemption_points_in_task_at_least_two
_ _ H_valid_model_with_fixed_preemption_points _ H_tsk_in_ts POSt) ⇒ Fact2.
(* ----------------------------------[ coqtop ]---------------------------------
2 subgoals (ID 705)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
Fact2 : 1 < size (task_preemption_points tsk)
============================
0 < size (distances (task_preemption_points tsk))
subgoal 2 (ID 621) is:
last0 (distances (task_preemption_points tsk)) <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
move: (Fact2) ⇒ Fact3.
(* ----------------------------------[ coqtop ]---------------------------------
2 subgoals (ID 707)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
Fact2, Fact3 : 1 < size (task_preemption_points tsk)
============================
0 < size (distances (task_preemption_points tsk))
subgoal 2 (ID 621) is:
last0 (distances (task_preemption_points tsk)) <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
by rewrite size_of_seq_of_distances // addn1 ltnS // in Fact2.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 621)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
last0 (distances (task_preemption_points tsk)) <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
- apply leq_trans with (task_max_nonpreemptive_segment tsk).
(* ----------------------------------[ coqtop ]---------------------------------
2 subgoals (ID 752)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
last0 (distances (task_preemption_points tsk)) <=
task_max_nonpreemptive_segment tsk
subgoal 2 (ID 753) is:
task_max_nonpreemptive_segment tsk <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
+ by apply last_of_seq_le_max_of_seq.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 753)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
task_max_nonpreemptive_segment tsk <= succn (task_cost tsk)
----------------------------------------------------------------------------- *)
+ rewrite -END; last by done.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 759)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
task_max_nonpreemptive_segment tsk <=
succn (last0 (task_preemption_points tsk))
----------------------------------------------------------------------------- *)
apply ltnW; rewrite ltnS; try done.
(* ----------------------------------[ coqtop ]---------------------------------
1 subgoal (ID 765)
Task : TaskType
H : TaskCost Task
Job : JobType
H0 : JobTask Job Task
H1 : JobArrival Job
H2 : JobCost Job
arr_seq : arrival_sequence Job
H_arrival_times_are_consistent : consistent_arrival_times arr_seq
H_arr_seq_is_a_set : arrival_sequence_uniq arr_seq
ts : seq Task
H_all_jobs_from_taskset : all_jobs_from_taskset arr_seq ts
H_job_cost_le_task_cost : cost_of_jobs_from_arrival_sequence_le_task_cost
arr_seq
H3 : JobPreemptionPoints Job
H4 : TaskPreemptionPoints Task
H_valid_model_with_fixed_preemption_points : valid_fixed_preemption_points_model
arr_seq ts
H5 : 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 (processor_state Job)
H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
arr_seq
H_valid_schedule_with_limited_preemptions : valid_schedule_with_limited_preemptions
arr_seq sched
H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
higher_eq_priority : FP_policy Task
H_priority_is_reflexive : reflexive_priorities
H_priority_is_transitive : transitive_priorities
H_sequential_tasks : sequential_tasks sched
H_work_conserving : work_conserving arr_seq sched
H_respects_policy : respects_policy_at_preemption_point arr_seq sched
task_rbf := task_request_bound_function tsk : duration -> nat
total_hep_rbf := total_hep_request_bound_function_FP higher_eq_priority ts
tsk : duration -> nat
total_ohep_rbf := total_ohep_request_bound_function_FP higher_eq_priority
ts tsk : duration -> nat
response_time_bounded_by := task_response_time_bound arr_seq sched
: Task -> duration -> Prop
blocking_bound := \max_(tsk_other <- ts | ~~
higher_eq_priority tsk_other tsk)
(task_max_nonpreemptive_segment tsk_other - ε) : nat
L : duration
H_L_positive : 0 < L
H_fixed_point : L = blocking_bound + total_hep_rbf L
is_in_search_space := fun A : duration =>
(A < L) && (task_rbf A != task_rbf (A + ε))
: duration -> bool
R : nat
H_R_is_maximum : forall A : duration,
is_in_search_space A ->
exists F : duration,
A + F =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + F) /\
F + (task_last_nonpr_segment tsk - ε) <= R
MLP : valid_limited_preemptions_job_model arr_seq
BEG : task_beginning_of_execution_in_preemption_points ts
END : task_end_of_execution_in_preemption_points ts
INCR : task_preemption_points_is_nondecreasing_sequence ts
HYP1 : job_consists_of_the_same_number_of_segments_as_task arr_seq
HYP2 : lengths_of_task_segments_bound_length_of_job_segments arr_seq
HYP3 : task_segments_are_nonempty ts
BEGj : beginning_of_execution_in_preemption_points arr_seq
ENDj : end_of_execution_in_preemption_points arr_seq
POSt : 0 < task_cost tsk
A : duration
SP : (A < L) &&
(task_request_bound_function tsk A
!= task_request_bound_function tsk (A + ε))
FF : duration
EQ1 : A + FF =
blocking_bound +
(task_rbf (A + ε) - (task_last_nonpr_segment tsk - ε)) +
total_ohep_rbf (A + FF)
EQ2 : FF + (task_last_nonpr_segment tsk - ε) <= R
============================
task_max_nonpreemptive_segment tsk <= last0 (task_preemption_points tsk)
----------------------------------------------------------------------------- *)
by apply max_distance_in_seq_le_last_element_of_seq; eauto 2.
(* ----------------------------------[ coqtop ]---------------------------------
No more subgoals.
----------------------------------------------------------------------------- *)
Qed.
End RTAforFixedPreemptionPointsModelwithArrivalCurves.