Library rt.restructuring.analysis.fixed_priority.rta.nonpr_reg.concrete_models.floating


(* ----------------------------------[ 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.floating.
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 Model with Floating Non-Preemptive Regions

In this module we prove the RTA theorem for floating non-preemptive regions FP model.
Consider any type of tasks ...
  Context {Task : TaskType}.
  Context `{TaskCost Task}.

... and any type of jobs associated with these tasks.
  Context {Job : JobType}.
  Context `{JobTask Job Task}.
  Context `{JobArrival Job}.
  Context `{JobCost Job}.

Consider any arrival sequence with consistent, non-duplicate arrivals.
Assume we have the model with floating nonpreemptive regions. I.e., for each task only the length of the maximal nonpreemptive segment is known _and_ each job level is divided into a number of nonpreemptive segments by inserting preemption points.
Consider an arbitrary task set ts, ...
  Variable ts : list Task.

... assume that all jobs come from the task set, ...
... and the cost of a job cannot be larger than the task cost.
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.
Let tsk be any task in ts that is to be analyzed.
  Variable tsk : Task.
  Hypothesis H_tsk_in_ts : tsk \in ts.

Next, consider any ideal uniprocessor schedule with limited preemptions of this arrival sequence ...
... where jobs do not execute before their arrival or after completion.
Consider an FP policy that indicates a higher-or-equal priority relation, and assume that the relation is reflexive and transitive.
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.
Next, we define a bound for the priority inversion caused by tasks of lower priority.
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.

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 : duration.
  Hypothesis H_R_is_maximum:
     (A : duration),
      is_in_search_space A
       (F : duration),
        A + F = blocking_bound + task_rbf (A + ε) + total_ohep_rbf (A + F)
        F 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 with floating nonpreemptive regions.
  Theorem uniprocessor_response_time_bound_fp_with_floating_nonpreemptive_regions:
    response_time_bounded_by tsk R.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 280)
  
  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
  H3 : JobPreemptionPoints Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H_valid_task_model_with_floating_nonpreemptive_regions : 
  valid_model_with_floating_nonpreemptive_regions 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
  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_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 : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     blocking_bound + task_rbf (A + ε) +
                     total_ohep_rbf (A + F) /\ F <= R
  ============================
  response_time_bounded_by tsk R

----------------------------------------------------------------------------- *)



  Proof.
    move: (H_valid_task_model_with_floating_nonpreemptive_regions) ⇒ [LIMJ JMLETM].

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 292)
  
  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
  H3 : JobPreemptionPoints Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H_valid_task_model_with_floating_nonpreemptive_regions : 
  valid_model_with_floating_nonpreemptive_regions 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
  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_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 : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     blocking_bound + task_rbf (A + ε) +
                     total_ohep_rbf (A + F) /\ F <= R
  LIMJ : valid_limited_preemptions_job_model arr_seq
  JMLETM : job_max_np_segment_le_task_max_np_segment arr_seq
  ============================
  response_time_bounded_by tsk R

----------------------------------------------------------------------------- *)


    move: (LIMJ) ⇒ [BEG [END _]].

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 315)
  
  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
  H3 : JobPreemptionPoints Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H_valid_task_model_with_floating_nonpreemptive_regions : 
  valid_model_with_floating_nonpreemptive_regions 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
  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_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 : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     blocking_bound + task_rbf (A + ε) +
                     total_ohep_rbf (A + F) /\ F <= R
  LIMJ : valid_limited_preemptions_job_model arr_seq
  JMLETM : job_max_np_segment_le_task_max_np_segment arr_seq
  BEG : beginning_of_execution_in_preemption_points arr_seq
  END : end_of_execution_in_preemption_points arr_seq
  ============================
  response_time_bounded_by tsk R

----------------------------------------------------------------------------- *)


    eapply uniprocessor_response_time_bound_fp_with_bounded_nonpreemptive_segments.

(* ----------------------------------[ coqtop ]---------------------------------

21 focused subgoals
(shelved: 2) (ID 331)
  
  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
  H3 : JobPreemptionPoints Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H_valid_task_model_with_floating_nonpreemptive_regions : 
  valid_model_with_floating_nonpreemptive_regions 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
  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_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 : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     blocking_bound + task_rbf (A + ε) +
                     total_ohep_rbf (A + F) /\ F <= R
  LIMJ : valid_limited_preemptions_job_model arr_seq
  JMLETM : job_max_np_segment_le_task_max_np_segment arr_seq
  BEG : beginning_of_execution_in_preemption_points arr_seq
  END : end_of_execution_in_preemption_points arr_seq
  ============================
  consistent_arrival_times arr_seq

subgoal 2 (ID 332) is:
 arrival_sequence_uniq arr_seq
subgoal 3 (ID 333) is:
 jobs_come_from_arrival_sequence sched arr_seq
subgoal 4 (ID 334) is:
 jobs_must_arrive_to_execute sched
subgoal 5 (ID 335) is:
 completed_jobs_dont_execute sched
subgoal 6 (ID 336) is:
 valid_model_with_bounded_nonpreemptive_segments arr_seq sched
subgoal 7 (ID 337) is:
 reflexive_priorities
subgoal 8 (ID 338) is:
 transitive_priorities
subgoal 9 (ID 339) is:
 sequential_tasks sched
subgoal 10 (ID 340) is:
 work_conserving arr_seq sched
subgoal 11 (ID 341) is:
 respects_policy_at_preemption_point arr_seq sched
subgoal 12 (ID 342) is:
 all_jobs_from_taskset arr_seq ?ts
subgoal 13 (ID 343) is:
 cost_of_jobs_from_arrival_sequence_le_task_cost arr_seq
subgoal 14 (ID 344) is:
 valid_taskset_arrival_curve ?ts max_arrivals
subgoal 15 (ID 345) is:
 taskset_respects_max_arrivals arr_seq ?ts
subgoal 16 (ID 346) is:
 tsk \in ?ts
subgoal 17 (ID 347) is:
 valid_preemption_model arr_seq sched
subgoal 18 (ID 348) is:
 valid_rtct.valid_task_run_to_completion_threshold arr_seq tsk
subgoal 19 (ID 349) is:
 0 < ?L
subgoal 20 (ID 350) is:
 ?L =
 response_time_bound.blocking_bound higher_eq_priority ?ts tsk +
 total_hep_request_bound_function_FP higher_eq_priority ?ts tsk ?L
subgoal 21 (ID 351) 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 ?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

----------------------------------------------------------------------------- *)


    all: eauto 2 with basic_facts.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 351)
  
  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
  H3 : JobPreemptionPoints Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H_valid_task_model_with_floating_nonpreemptive_regions : 
  valid_model_with_floating_nonpreemptive_regions 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
  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_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 : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     blocking_bound + task_rbf (A + ε) +
                     total_ohep_rbf (A + F) /\ F <= R
  LIMJ : valid_limited_preemptions_job_model arr_seq
  JMLETM : job_max_np_segment_le_task_max_np_segment arr_seq
  BEG : beginning_of_execution_in_preemption_points arr_seq
  END : end_of_execution_in_preemption_points arr_seq
  ============================
  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 383)
  
  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
  H3 : JobPreemptionPoints Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H_valid_task_model_with_floating_nonpreemptive_regions : 
  valid_model_with_floating_nonpreemptive_regions 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
  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_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 : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     blocking_bound + task_rbf (A + ε) +
                     total_ohep_rbf (A + F) /\ F <= R
  LIMJ : valid_limited_preemptions_job_model arr_seq
  JMLETM : job_max_np_segment_le_task_max_np_segment arr_seq
  BEG : beginning_of_execution_in_preemption_points arr_seq
  END : end_of_execution_in_preemption_points arr_seq
  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

----------------------------------------------------------------------------- *)


    rewrite subnn subn0.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 393)
  
  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
  H3 : JobPreemptionPoints Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H_valid_task_model_with_floating_nonpreemptive_regions : 
  valid_model_with_floating_nonpreemptive_regions 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
  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_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 : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     blocking_bound + task_rbf (A + ε) +
                     total_ohep_rbf (A + F) /\ F <= R
  LIMJ : valid_limited_preemptions_job_model arr_seq
  JMLETM : job_max_np_segment_le_task_max_np_segment arr_seq
  BEG : beginning_of_execution_in_preemption_points arr_seq
  END : end_of_execution_in_preemption_points arr_seq
  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 + ε) +
    total_ohep_request_bound_function_FP higher_eq_priority ts tsk (A + F) /\
    F + 0 <= R

----------------------------------------------------------------------------- *)


    destruct (H_R_is_maximum _ SP) as [F [EQ LE]].

(* ----------------------------------[ 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
  H3 : JobPreemptionPoints Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H_valid_task_model_with_floating_nonpreemptive_regions : 
  valid_model_with_floating_nonpreemptive_regions 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
  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_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 : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     blocking_bound + task_rbf (A + ε) +
                     total_ohep_rbf (A + F) /\ F <= R
  LIMJ : valid_limited_preemptions_job_model arr_seq
  JMLETM : job_max_np_segment_le_task_max_np_segment arr_seq
  BEG : beginning_of_execution_in_preemption_points arr_seq
  END : end_of_execution_in_preemption_points arr_seq
  A : duration
  SP : (A < L) &&
       (task_request_bound_function tsk A
        != task_request_bound_function tsk (A + ε))
  F : duration
  EQ : A + F = blocking_bound + task_rbf (A + ε) + total_ohep_rbf (A + F)
  LE : F <= R
  ============================
  exists F0 : duration,
    A + F0 =
    response_time_bound.blocking_bound higher_eq_priority ts tsk +
    task_request_bound_function tsk (A + ε) +
    total_ohep_request_bound_function_FP higher_eq_priority ts tsk (A + F0) /\
    F0 + 0 <= R

----------------------------------------------------------------------------- *)


      by F; rewrite addn0; split.

(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


  Qed.

End RTAforFloatingModelwithArrivalCurves.