Library prosa.results.fixed_priority.rta.bounded_pi


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

Welcome to Coq 8.11.2 (June 2020)

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


Require Export prosa.model.schedule.priority_driven.
Require Export prosa.analysis.facts.busy_interval.busy_interval.
Require Import prosa.analysis.abstract.ideal_jlfp_rta.

Throughout this file, we assume ideal uni-processor schedules.
Require Import prosa.model.processor.ideal.

Throughout this file, we assume the basic (i.e., Liu & Layland) readiness model.
Require Import prosa.model.readiness.basic.

Abstract RTA for FP-schedulers with Bounded Priority Inversion

In this module we instantiate the Abstract Response-Time analysis (aRTA) to FP-schedulers for ideal uni-processor model of real-time tasks with arbitrary arrival models.
Given FP priority policy and an ideal uni-processor scheduler model, we can explicitly specify [interference], [interfering_workload], and [interference_bound_function]. In this settings, we can define natural notions of service, workload, busy interval, etc. The important feature of this instantiation is that we can induce the meaningful notion of priority inversion. However, we do not specify the exact cause of priority inversion (as there may be different reasons for this, like execution of a non-preemptive segment or blocking due to resource locking). We only assume that that a priority inversion is bounded.
Consider any type of tasks ...
  Context {Task : TaskType}.
  Context `{TaskCost Task}.
  Context `{TaskRunToCompletionThreshold Task}.

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

Consider any arrival sequence with consistent, non-duplicate arrivals.
Next, consider any ideal uniprocessor schedule of this arrival sequence ...
... where jobs do not execute before their arrival or after completion.
Note that we differentiate between abstract and classical notions of work conserving schedule.
We assume that the schedule is a work-conserving schedule in the _classical_ sense, and later prove that the hypothesis about abstract work-conservation also holds.
Assume we have sequential tasks, i.e, jobs from the same task execute in the order of their arrival.
Assume that a job cost cannot be larger than a task cost.
Consider an arbitrary task set ts.
  Variable ts : list Task.

Next, we assume that all jobs come from the task set.
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.

Consider a valid preemption model...
...and a valid task run-to-completion threshold function. That is, [task_run_to_completion_threshold tsk] is (1) no bigger than [tsk]'s cost, (2) for any job of task [tsk] job_run_to_completion_threshold is bounded by task_run_to_completion_threshold.
Consider an FP policy that indicates a higher-or-equal priority relation, and assume that the relation is reflexive. Note that we do not relate the FP policy with the scheduler. However, we define functions for Interference and Interfering Workload that actively use the concept of priorities. We require the FP policy to be reflexive, so a job cannot cause lower-priority interference (i.e. priority inversion) to itself.
  Context `{FP_policy Task}.
  Hypothesis H_priority_is_reflexive : reflexive_priorities.

For clarity, let's define some local names.
We introduce [task_rbf] as an abbreviation of the task request bound function, which is defined as [task_cost(tsk) × max_arrivals(tsk,Δ)].
Using the sum of individual request bound functions, we define the request bound function of all tasks with higher-or-equal priority (with respect to [tsk]).
Similarly, we define the request bound function of all tasks other than [tsk] with higher-or-equal priority (with respect to [tsk]).
Assume that there exists a constant priority_inversion_bound that bounds the length of any priority inversion experienced by any job of [tsk]. Since we analyze only task [tsk], we ignore the lengths of priority inversions incurred by any other tasks.
Let L be any positive fixed point of the busy interval recurrence.
  Variable L : duration.
  Hypothesis H_L_positive : L > 0.
  Hypothesis H_fixed_point : L = priority_inversion_bound + total_hep_rbf L.

To reduce the time complexity of the analysis, recall the notion of search space. Intuitively, this corresponds to all "interesting" arrival offsets that the job under analysis might have with regard to the beginning of its busy-window.
  Definition is_in_search_space A := (A < L) && (task_rbf A != task_rbf (A + ε)).

Let R be a value that upper-bounds the solution of each response-time recurrence, i.e., for any relative arrival time A in the search space, there exists a corresponding solution F such that [F + (task cost - task lock-in service) <= R].
  Variable R : duration.
  Hypothesis H_R_is_maximum :
     (A : duration),
      is_in_search_space A
       (F : duration),
        A + F = priority_inversion_bound
                + (task_rbf (A + ε) - (task_cost tsk - task_run_to_completion_threshold tsk))
                + total_ohep_rbf (A + F)
        F + (task_cost tsk - task_run_to_completion_threshold tsk) R.

Instantiation of Interference We say that job j incurs interference at time t iff it cannot execute due to a higher-or-equal-priority job being scheduled, or if it incurs a priority inversion.
  Let interference (j : Job) (t : instant) :=
    ideal_jlfp_rta.interference sched j t.

Instantiation of Interfering Workload The interfering workload, in turn, is defined as the sum of the priority inversion function and interfering workload of jobs with higher or equal priority.
Finally, we define the interference bound function as the sum of the priority interference bound and the higher-or-equal-priority workload.

Filling Out Hypotheses Of Abstract RTA Theorem

In this section we prove that all preconditions necessary to use the abstract theorem are satisfied.
First, we prove that in the instantiation of interference and interfering workload, we really take into account everything that can interfere with [tsk]'s jobs, and thus, the scheduler satisfies the abstract notion of work conserving schedule.
    Lemma instantiated_i_and_w_are_consistent_with_schedule:
      work_conserving_ab tsk interference interfering_workload.

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

1 subgoal (ID 1366)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  ============================
  work_conserving_ab tsk interference interfering_workload

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


    Proof.
      intros j t1 t2 t ARR TSK POS BUSY NEQ; split; intros HYP.

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

2 subgoals (ID 1380)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  HYP : ~ interference j t
  ============================
  scheduled_at sched j t

subgoal 2 (ID 1381) is:
 ~ interference j t

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


      - move: HYP ⇒ /negP.

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

2 subgoals (ID 1427)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  ============================
  ~~ interference j t -> scheduled_at sched j t

subgoal 2 (ID 1381) is:
 ~ interference j t

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


        rewrite negb_or /is_priority_inversion /is_priority_inversion
                /is_interference_from_another_hep_job.

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

2 subgoals (ID 1440)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  ============================
  ~~ match sched t with
     | Some jlp => ~~ hep_job jlp j
     | None => false
     end &&
  ~~
  match sched t with
  | Some jhp => hep_job jhp j && (jhp != j)
  | None => false
  end -> scheduled_at sched j t

subgoal 2 (ID 1381) is:
 ~ interference j t

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


        move ⇒ /andP [HYP1 HYP2].

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

2 subgoals (ID 1480)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  HYP1 : ~~
         match sched t with
         | Some jlp => ~~ hep_job jlp j
         | None => false
         end
  HYP2 : ~~
         match sched t with
         | Some jhp => hep_job jhp j && (jhp != j)
         | None => false
         end
  ============================
  scheduled_at sched j t

subgoal 2 (ID 1381) is:
 ~ interference j t

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


        case SCHED: (sched t) ⇒ [s | ].

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

3 subgoals (ID 1553)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  HYP1 : ~~
         match sched t with
         | Some jlp => ~~ hep_job jlp j
         | None => false
         end
  HYP2 : ~~
         match sched t with
         | Some jhp => hep_job jhp j && (jhp != j)
         | None => false
         end
  s : Job
  SCHED : sched t = Some s
  ============================
  scheduled_at sched j t

subgoal 2 (ID 1615) is:
 scheduled_at sched j t
subgoal 3 (ID 1381) is:
 ~ interference j t

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


        + rewrite SCHED in HYP1, HYP2.

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

3 subgoals (ID 1687)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  s : Job
  SCHED : sched t = Some s
  HYP1 : ~~ ~~ hep_job s j
  HYP2 : ~~ (hep_job s j && (s != j))
  ============================
  scheduled_at sched j t

subgoal 2 (ID 1615) is:
 scheduled_at sched j t
subgoal 3 (ID 1381) is:
 ~ interference j t

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


          move: HYP1 HYP2.
(* ----------------------------------[ coqtop ]---------------------------------

3 subgoals (ID 1691)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  s : Job
  SCHED : sched t = Some s
  ============================
  ~~ ~~ hep_job s j -> ~~ (hep_job s j && (s != j)) -> scheduled_at sched j t

subgoal 2 (ID 1615) is:
 scheduled_at sched j t
subgoal 3 (ID 1381) is:
 ~ interference j t

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


          rewrite !Bool.negb_involutive negb_and Bool.negb_involutive.

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

3 subgoals (ID 1702)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  s : Job
  SCHED : sched t = Some s
  ============================
  hep_job s j -> ~~ hep_job s j || (s == j) -> scheduled_at sched j t

subgoal 2 (ID 1615) is:
 scheduled_at sched j t
subgoal 3 (ID 1381) is:
 ~ interference j t

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


          moveHYP1 /orP [/negP HYP2| /eqP HYP2].

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

4 subgoals (ID 1802)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  s : Job
  SCHED : sched t = Some s
  HYP1 : hep_job s j
  HYP2 : ~ hep_job s j
  ============================
  scheduled_at sched j t

subgoal 2 (ID 1803) is:
 scheduled_at sched j t
subgoal 3 (ID 1615) is:
 scheduled_at sched j t
subgoal 4 (ID 1381) is:
 ~ interference j t

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


          × by exfalso.

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

3 subgoals (ID 1803)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  s : Job
  SCHED : sched t = Some s
  HYP1 : hep_job s j
  HYP2 : s = j
  ============================
  scheduled_at sched j t

subgoal 2 (ID 1615) is:
 scheduled_at sched j t
subgoal 3 (ID 1381) is:
 ~ interference j t

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


          × by subst s; rewrite scheduled_at_def //; apply eqprop_to_eqbool.

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

2 subgoals (ID 1615)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  HYP1 : ~~
         match sched t with
         | Some jlp => ~~ hep_job jlp j
         | None => false
         end
  HYP2 : ~~
         match sched t with
         | Some jhp => hep_job jhp j && (jhp != j)
         | None => false
         end
  SCHED : sched t = None
  ============================
  scheduled_at sched j t

subgoal 2 (ID 1381) is:
 ~ interference j t

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


        + exfalso; clear HYP1 HYP2.

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

2 subgoals (ID 1872)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  SCHED : sched t = None
  ============================
  False

subgoal 2 (ID 1381) is:
 ~ interference j t

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


          eapply instantiated_busy_interval_equivalent_edf_busy_interval in BUSY; eauto with basic_facts.

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

2 subgoals (ID 1882)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  NEQ : t1 <= t < t2
  SCHED : sched t = None
  ============================
  False

subgoal 2 (ID 1381) is:
 ~ interference j t

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


            by move: BUSY ⇒ [PREF _]; eapply not_quiet_implies_not_idle;
                              eauto 2 using eqprop_to_eqbool.

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

1 subgoal (ID 1381)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  HYP : scheduled_at sched j t
  ============================
  ~ interference j t

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


      - move: (HYP); rewrite scheduled_at_def; move ⇒ /eqP HYP2; apply/negP.

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

1 focused subgoal
(shelved: 1) (ID 2148)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  HYP : scheduled_at sched j t
  HYP2 : sched t = Some j
  ============================
  ~~ interference j t

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


        rewrite /interference /ideal_jlfp_rta.interference /is_priority_inversion
                  /is_interference_from_another_hep_job HYP2 negb_or.

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

1 subgoal (ID 2164)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  HYP : scheduled_at sched j t
  HYP2 : sched t = Some j
  ============================
  ~~ ~~ hep_job j j && ~~ (hep_job j j && (j != j))

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


         apply/andP; split.

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

2 subgoals (ID 2190)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  HYP : scheduled_at sched j t
  HYP2 : sched t = Some j
  ============================
  ~~ ~~ hep_job j j

subgoal 2 (ID 2191) is:
 ~~ (hep_job j j && (j != j))

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


         + rewrite Bool.negb_involutive; eauto 2.

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

2 subgoals (ID 2194)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  HYP : scheduled_at sched j t
  HYP2 : sched t = Some j
  ============================
  hep_job j j

subgoal 2 (ID 2191) is:
 ~~ (hep_job j j && (j != j))

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


             by eapply H_priority_is_reflexive with (t := 0).

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

1 subgoal (ID 2191)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  t : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  NEQ : t1 <= t < t2
  HYP : scheduled_at sched j t
  HYP2 : sched t = Some j
  ============================
  ~~ (hep_job j j && (j != j))

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


         + by rewrite negb_and Bool.negb_involutive; apply/orP; right.

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

No more subgoals.

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


    Qed.

Next, we prove that the interference and interfering workload functions are consistent with sequential tasks.
    Lemma instantiated_interference_and_workload_consistent_with_sequential_tasks:
      interference_and_workload_consistent_with_sequential_tasks
        arr_seq sched tsk interference interfering_workload.

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

1 subgoal (ID 1372)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  ============================
  interference_and_workload_consistent_with_sequential_tasks arr_seq sched
    tsk interference interfering_workload

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


    Proof.
      intros j t1 t2 ARR TSK POS BUSY.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1380)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  ============================
  task_workload_between arr_seq tsk 0 t1 =
  task_service_of_jobs_in sched tsk (arrivals_between arr_seq 0 t1) 0 t1

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


      eapply instantiated_busy_interval_equivalent_edf_busy_interval in BUSY; eauto with basic_facts.

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

1 subgoal (ID 1390)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  ============================
  task_workload_between arr_seq tsk 0 t1 =
  task_service_of_jobs_in sched tsk (arrivals_between arr_seq 0 t1) 0 t1

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


      eapply all_jobs_have_completed_equiv_workload_eq_service; eauto 2; intros s ARRs TSKs.

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

1 subgoal (ID 1487)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  s : Job
  ARRs : s \in arrivals_between arr_seq 0 t1
  TSKs : job_of_task tsk s
  ============================
  completed_by sched s t1

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


      move: (BUSY) ⇒ [[_ [QT _]] _].

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

1 subgoal (ID 1520)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  s : Job
  ARRs : s \in arrivals_between arr_seq 0 t1
  TSKs : job_of_task tsk s
  QT : busy_interval.quiet_time arr_seq sched j t1
  ============================
  completed_by sched s t1

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


      apply QT.

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

3 subgoals (ID 1521)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  s : Job
  ARRs : s \in arrivals_between arr_seq 0 t1
  TSKs : job_of_task tsk s
  QT : busy_interval.quiet_time arr_seq sched j t1
  ============================
  arrives_in arr_seq s

subgoal 2 (ID 1522) is:
 hep_job s j
subgoal 3 (ID 1523) is:
 arrived_before s t1

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


      - by apply in_arrivals_implies_arrived in ARRs.

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

2 subgoals (ID 1522)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  s : Job
  ARRs : s \in arrivals_between arr_seq 0 t1
  TSKs : job_of_task tsk s
  QT : busy_interval.quiet_time arr_seq sched j t1
  ============================
  hep_job s j

subgoal 2 (ID 1523) is:
 arrived_before s t1

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


      - move: TSKs ⇒ /eqP TSKs.

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

2 subgoals (ID 1562)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  s : Job
  ARRs : s \in arrivals_between arr_seq 0 t1
  QT : busy_interval.quiet_time arr_seq sched j t1
  TSKs : job_task s = tsk
  ============================
  hep_job s j

subgoal 2 (ID 1523) is:
 arrived_before s t1

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


        rewrite /hep_job /FP_to_JLFP TSK -TSKs; eauto 2.

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

2 subgoals (ID 1572)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  s : Job
  ARRs : s \in arrivals_between arr_seq 0 t1
  QT : busy_interval.quiet_time arr_seq sched j t1
  TSKs : job_task s = tsk
  ============================
  hep_task (job_task s) (job_task s)

subgoal 2 (ID 1523) is:
 arrived_before s t1

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


          by eapply (H_priority_is_reflexive 0); eauto.

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

1 subgoal (ID 1523)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  t1, t2 : instant
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  s : Job
  ARRs : s \in arrivals_between arr_seq 0 t1
  TSKs : job_of_task tsk s
  QT : busy_interval.quiet_time arr_seq sched j t1
  ============================
  arrived_before s t1

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


      - by eapply in_arrivals_implies_arrived_before; eauto 2.

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

No more subgoals.

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


    Qed.

Recall that L is assumed to be a fixed point of the busy interval recurrence. Thanks to this fact, we can prove that every busy interval (according to the concrete definition) is bounded. In addition, we know that the conventional concept of busy interval and the one obtained from the abstract definition (with the interference and interfering workload) coincide. Thus, it follows that any busy interval (in the abstract sense) is bounded.
    Lemma instantiated_busy_intervals_are_bounded:
      busy_intervals_are_bounded_by arr_seq sched tsk interference interfering_workload L.

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

1 subgoal (ID 1380)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  ============================
  busy_intervals_are_bounded_by arr_seq sched tsk interference
    interfering_workload L

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


    Proof.
      intros j ARR TSK POS.

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

1 subgoal (ID 1385)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  ============================
  exists t1 t2 : nat,
    t1 <= job_arrival j < t2 /\
    t2 <= t1 + L /\
    busy_interval sched interference interfering_workload j t1 t2

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


      edestruct (exists_busy_interval) with (delta := L) as [t1 [t2 [T1 [T2 GGG]]]]; eauto 2.

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

2 subgoals (ID 1421)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  ============================
  forall t : instant,
  priority_inversion_bound +
  workload_of_higher_or_equal_priority_jobs j
    (arrivals_between arr_seq t (t + L)) <= L

subgoal 2 (ID 1440) is:
 exists t0 t3 : nat,
   t0 <= job_arrival j < t3 /\
   t3 <= t0 + L /\
   busy_interval sched interference interfering_workload j t0 t3

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


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

1 subgoal (ID 1421)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  ============================
  forall t : instant,
  priority_inversion_bound +
  workload_of_higher_or_equal_priority_jobs j
    (arrivals_between arr_seq t (t + L)) <= L

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


by intros; rewrite {2}H_fixed_point leq_add //; apply total_workload_le_total_rbf'.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1440)

subgoal 1 (ID 1440) is:
 exists t0 t3 : nat,
   t0 <= job_arrival j < t3 /\
   t3 <= t0 + L /\
   busy_interval sched interference interfering_workload j t0 t3

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


}

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

1 subgoal (ID 1440)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  t1, t2 : nat
  T1 : t1 <= job_arrival j < t2
  T2 : t2 <= t1 + L
  GGG : busy_interval.busy_interval arr_seq sched j t1 t2
  ============================
  exists t0 t3 : nat,
    t0 <= job_arrival j < t3 /\
    t3 <= t0 + L /\
    busy_interval sched interference interfering_workload j t0 t3

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


       t1, t2; split; first by done.

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

1 subgoal (ID 1538)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  t1, t2 : nat
  T1 : t1 <= job_arrival j < t2
  T2 : t2 <= t1 + L
  GGG : busy_interval.busy_interval arr_seq sched j t1 t2
  ============================
  t2 <= t1 + L /\
  busy_interval sched interference interfering_workload j t1 t2

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


      split.

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

2 subgoals (ID 1540)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  t1, t2 : nat
  T1 : t1 <= job_arrival j < t2
  T2 : t2 <= t1 + L
  GGG : busy_interval.busy_interval arr_seq sched j t1 t2
  ============================
  t2 <= t1 + L

subgoal 2 (ID 1541) is:
 busy_interval sched interference interfering_workload j t1 t2

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


      - by done.

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

1 subgoal (ID 1541)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  POS : 0 < job_cost j
  t1, t2 : nat
  T1 : t1 <= job_arrival j < t2
  T2 : t2 <= t1 + L
  GGG : busy_interval.busy_interval arr_seq sched j t1 t2
  ============================
  busy_interval sched interference interfering_workload j t1 t2

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


      - by eapply instantiated_busy_interval_equivalent_edf_busy_interval; eauto 2 with basic_facts.

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

No more subgoals.

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


    Qed.

Next, we prove that IBF is indeed an interference bound.
Recall that in module abstract_seq_RTA hypothesis task_interference_is_bounded_by expects to receive a function that maps some task t, the relative arrival time of a job j of task t, and the length of the interval to the maximum amount of interference (for more details see files limited.abstract_RTA.definitions and limited.abstract_RTA.abstract_seq_rta).
However, in this module we analyze only one task -- [tsk], therefore it is “hard-coded” inside the interference bound function IBF. Moreover, in case of a model with fixed priorities, interference that some job j incurs from higher-or-equal priority jobs does not depend on the relative arrival time of job j. Therefore, in order for the IBF signature to match the required signature in module abstract_seq_RTA, we wrap the IBF function in a function that accepts, but simply ignores, the task and the relative arrival time.
    Lemma instantiated_task_interference_is_bounded:
      task_interference_is_bounded_by
        arr_seq sched tsk interference interfering_workload (fun t A RIBF R).

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

1 subgoal (ID 1389)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  ============================
  task_interference_is_bounded_by arr_seq sched tsk interference
    interfering_workload (fun=> (fun=> [eta IBF]))

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


    Proof.
      intros ? ? ? ? ARR TSK ? NCOMPL BUSY; simpl.

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

1 subgoal (ID 1401)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  ============================
  cumul_task_interference arr_seq sched interference tsk t2 t1 (t1 + R0) <=
  IBF R0

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


      move: (posnP (@job_cost _ H3 j)) ⇒ [ZERO|POS].

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

2 subgoals (ID 1418)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  ZERO : job_cost j = 0
  ============================
  cumul_task_interference arr_seq sched interference tsk t2 t1 (t1 + R0) <=
  IBF R0

subgoal 2 (ID 1419) is:
 cumul_task_interference arr_seq sched interference tsk t2 t1 (t1 + R0) <=
 IBF R0

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


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

1 subgoal (ID 1418)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  ZERO : job_cost j = 0
  ============================
  cumul_task_interference arr_seq sched interference tsk t2 t1 (t1 + R0) <=
  IBF R0

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


by exfalso; rewrite /completed_by ZERO in NCOMPL.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1419)

subgoal 1 (ID 1419) is:
 cumul_task_interference arr_seq sched interference tsk t2 t1 (t1 + R0) <=
 IBF R0

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


}

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

1 subgoal (ID 1419)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval sched interference interfering_workload j t1 t2
  POS : 0 < job_cost j
  ============================
  cumul_task_interference arr_seq sched interference tsk t2 t1 (t1 + R0) <=
  IBF R0

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


      eapply instantiated_busy_interval_equivalent_edf_busy_interval in BUSY; eauto 2 with basic_facts.

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

1 subgoal (ID 1532)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  ============================
  cumul_task_interference arr_seq sched interference tsk t2 t1 (t1 + R0) <=
  IBF R0

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


      rewrite /interference; erewrite cumulative_task_interference_split; eauto 2 with basic_facts; last first.

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

2 subgoals (ID 1569)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  ============================
  j \in arrivals_before arr_seq t2

subgoal 2 (ID 1565) is:
 cumulative_priority_inversion sched j t1 (t1 + R0) +
 cumulative_interference_from_hep_jobs_from_other_tasks sched j t1 (t1 + R0) <=
 IBF R0

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


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

1 subgoal (ID 1569)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  ============================
  j \in arrivals_before arr_seq t2

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


move: BUSY ⇒ [[_ [_ [_ /andP [GE LT]]]] _].

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

1 subgoal (ID 1704)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  POS : 0 < job_cost j
  GE : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  j \in arrivals_before arr_seq t2

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


          by eapply arrived_between_implies_in_arrivals; eauto 2.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1565)

subgoal 1 (ID 1565) is:
 cumulative_priority_inversion sched j t1 (t1 + R0) +
 cumulative_interference_from_hep_jobs_from_other_tasks sched j t1 (t1 + R0) <=
 IBF R0

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


}

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

1 subgoal (ID 1565)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  ============================
  cumulative_priority_inversion sched j t1 (t1 + R0) +
  cumulative_interference_from_hep_jobs_from_other_tasks sched j t1 (t1 + R0) <=
  IBF R0

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


      unfold IBF, interference.

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

1 subgoal (ID 1715)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  ============================
  cumulative_priority_inversion sched j t1 (t1 + R0) +
  cumulative_interference_from_hep_jobs_from_other_tasks sched j t1 (t1 + R0) <=
  priority_inversion_bound + total_ohep_rbf R0

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


      rewrite leq_add; try done.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1724)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  ============================
  cumulative_priority_inversion sched j t1 (t1 + R0) <=
  priority_inversion_bound

subgoal 2 (ID 1725) is:
 cumulative_interference_from_hep_jobs_from_other_tasks sched j t1 (t1 + R0) <=
 total_ohep_rbf R0

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


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

1 subgoal (ID 1724)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  ============================
  cumulative_priority_inversion sched j t1 (t1 + R0) <=
  priority_inversion_bound

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


move: (H_priority_inversion_is_bounded j ARR TSK) ⇒ BOUND.

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

1 subgoal (ID 1771)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  BOUND : 0 < job_cost j ->
          priority_inversion_of_job_is_bounded_by arr_seq sched j
            priority_inversion_bound
  ============================
  cumulative_priority_inversion sched j t1 (t1 + R0) <=
  priority_inversion_bound

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


        apply leq_trans with (cumulative_priority_inversion sched j t1 (t1 + R0)); first by done.

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

1 subgoal (ID 1778)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  BOUND : 0 < job_cost j ->
          priority_inversion_of_job_is_bounded_by arr_seq sched j
            priority_inversion_bound
  ============================
  cumulative_priority_inversion sched j t1 (t1 + R0) <=
  priority_inversion_bound

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


        apply leq_trans with (cumulative_priority_inversion sched j t1 t2); last first.

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

2 subgoals (ID 1785)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  BOUND : 0 < job_cost j ->
          priority_inversion_of_job_is_bounded_by arr_seq sched j
            priority_inversion_bound
  ============================
  cumulative_priority_inversion sched j t1 t2 <= priority_inversion_bound

subgoal 2 (ID 1784) is:
 cumulative_priority_inversion sched j t1 (t1 + R0) <=
 cumulative_priority_inversion sched j t1 t2

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


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

1 subgoal (ID 1785)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  BOUND : 0 < job_cost j ->
          priority_inversion_of_job_is_bounded_by arr_seq sched j
            priority_inversion_bound
  ============================
  cumulative_priority_inversion sched j t1 t2 <= priority_inversion_bound

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


by apply BOUND; move: BUSY ⇒ [PREF QT2].
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1784)

subgoal 1 (ID 1784) is:
 cumulative_priority_inversion sched j t1 (t1 + R0) <=
 cumulative_priority_inversion sched j t1 t2
subgoal 2 (ID 1725) is:
 cumulative_interference_from_hep_jobs_from_other_tasks sched j t1 (t1 + R0) <=
 total_ohep_rbf R0

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


}

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

1 subgoal (ID 1784)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  BOUND : 0 < job_cost j ->
          priority_inversion_of_job_is_bounded_by arr_seq sched j
            priority_inversion_bound
  ============================
  cumulative_priority_inversion sched j t1 (t1 + R0) <=
  cumulative_priority_inversion sched j t1 t2

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


        rewrite [X in _ X](@big_cat_nat _ _ _ (t1 + R0)) //=.

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

3 subgoals (ID 1841)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  BOUND : 0 < job_cost j ->
          priority_inversion_of_job_is_bounded_by arr_seq sched j
            priority_inversion_bound
  ============================
  cumulative_priority_inversion sched j t1 (t1 + R0) <=
  \sum_(t1 <= i < t1 + R0) is_priority_inversion sched j i +
  \sum_(t1 + R0 <= i < t2) is_priority_inversion sched j i

subgoal 2 (ID 1864) is:
 t1 <= t1 + R0
subgoal 3 (ID 1887) is:
 t1 + R0 <= t2

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


        - by rewrite leq_addr.

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

2 subgoals (ID 1864)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  BOUND : 0 < job_cost j ->
          priority_inversion_of_job_is_bounded_by arr_seq sched j
            priority_inversion_bound
  ============================
  t1 <= t1 + R0

subgoal 2 (ID 1887) is:
 t1 + R0 <= t2

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


        - by rewrite leq_addr.

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

1 subgoal (ID 1887)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  BOUND : 0 < job_cost j ->
          priority_inversion_of_job_is_bounded_by arr_seq sched j
            priority_inversion_bound
  ============================
  t1 + R0 <= t2

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


        - by rewrite ltnW.

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

1 subgoal (ID 1725)

subgoal 1 (ID 1725) is:
 cumulative_interference_from_hep_jobs_from_other_tasks sched j t1 (t1 + R0) <=
 total_ohep_rbf R0

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


      }

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

1 subgoal (ID 1725)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  ============================
  cumulative_interference_from_hep_jobs_from_other_tasks sched j t1 (t1 + R0) <=
  total_ohep_rbf R0

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


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

1 subgoal (ID 1725)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  ============================
  cumulative_interference_from_hep_jobs_from_other_tasks sched j t1 (t1 + R0) <=
  total_ohep_rbf R0

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


erewrite instantiated_cumulative_interference_of_hep_tasks_equal_total_interference_of_hep_tasks;
          eauto 2; last by unfold quiet_time; move: BUSY ⇒ [[_ [T1 T2]] _].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1930)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  ============================
  service_of_jobs sched
    (fun jhp : Job => hep_job jhp j && (job_task jhp != job_task j))
    (arrivals_between arr_seq t1 (t1 + R0)) t1 (t1 + R0) <= 
  total_ohep_rbf R0

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


        apply leq_trans with
            (workload_of_jobs
               (fun jhp : Job ⇒ (FP_to_JLFP _ _) jhp j && (job_task jhp != job_task j))
               (arrivals_between arr_seq t1 (t1 + R0))
            ).

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

2 subgoals (ID 2035)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  ============================
  service_of_jobs sched
    (fun jhp : Job => hep_job jhp j && (job_task jhp != job_task j))
    (arrivals_between arr_seq t1 (t1 + R0)) t1 (t1 + R0) <=
  workload_of_jobs
    (fun jhp : Job =>
     FP_to_JLFP Job Task jhp j && (job_task jhp != job_task j))
    (arrivals_between arr_seq t1 (t1 + R0))

subgoal 2 (ID 2036) is:
 workload_of_jobs
   (fun jhp : Job =>
    FP_to_JLFP Job Task jhp j && (job_task jhp != job_task j))
   (arrivals_between arr_seq t1 (t1 + R0)) <= total_ohep_rbf R0

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


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

1 subgoal (ID 2035)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  ============================
  service_of_jobs sched
    (fun jhp : Job => hep_job jhp j && (job_task jhp != job_task j))
    (arrivals_between arr_seq t1 (t1 + R0)) t1 (t1 + R0) <=
  workload_of_jobs
    (fun jhp : Job =>
     FP_to_JLFP Job Task jhp j && (job_task jhp != job_task j))
    (arrivals_between arr_seq t1 (t1 + R0))

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


by apply service_of_jobs_le_workload; last apply ideal_proc_model_provides_unit_service.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2036)

subgoal 1 (ID 2036) is:
 workload_of_jobs
   (fun jhp : Job =>
    FP_to_JLFP Job Task jhp j && (job_task jhp != job_task j))
   (arrivals_between arr_seq t1 (t1 + R0)) <= total_ohep_rbf R0

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


 }

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

1 subgoal (ID 2036)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  ============================
  workload_of_jobs
    (fun jhp : Job =>
     FP_to_JLFP Job Task jhp j && (job_task jhp != job_task j))
    (arrivals_between arr_seq t1 (t1 + R0)) <= total_ohep_rbf R0

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


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

1 subgoal (ID 2036)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  ============================
  workload_of_jobs
    (fun jhp : Job =>
     FP_to_JLFP Job Task jhp j && (job_task jhp != job_task j))
    (arrivals_between arr_seq t1 (t1 + R0)) <= total_ohep_rbf R0

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


rewrite /workload_of_jobs /total_ohep_rbf /total_ohep_request_bound_function_FP.

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

1 subgoal (ID 2058)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  R0, t1, t2 : nat
  ARR : arrives_in arr_seq j
  TSK : job_task j = tsk
  H7 : t1 + R0 < t2
  NCOMPL : ~~ completed_by sched j (t1 + R0)
  BUSY : busy_interval.busy_interval arr_seq sched j t1 t2
  POS : 0 < job_cost j
  ============================
  \sum_(j0 <- arrivals_between arr_seq t1 (t1 + R0) | 
  FP_to_JLFP Job Task j0 j && (job_task j0 != job_task j)) 
  job_cost j0 <=
  \sum_(tsk_other <- ts | hep_task tsk_other tsk && (tsk_other != tsk))
     task_request_bound_function tsk_other R0

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


            by rewrite -TSK; apply total_workload_le_total_rbf.

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

No more subgoals.

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


        }

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

No more subgoals.

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


      }

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

No more subgoals.

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


    Qed.

Finally, we show that there exists a solution for the response-time recurrence.
Consider any job [j] of [tsk].
      Variable j : Job.
      Hypothesis H_j_arrives : arrives_in arr_seq j.
      Hypothesis H_job_of_tsk : job_task j = tsk.
      Hypothesis H_job_cost_positive: job_cost_positive j.

Given any job j of task [tsk] that arrives exactly A units after the beginning of the busy interval, the bound of the total interference incurred by j within an interval of length Δ is equal to [task_rbf (A + ε) - task_cost tsk + IBF Δ].
      Let total_interference_bound tsk A Δ :=
        task_rbf (A + ε) - task_cost tsk + IBF Δ.

Next, consider any A from the search space (in the abstract sense).
      Variable A : duration.
      Hypothesis H_A_is_in_abstract_search_space :
        search_space.is_in_search_space tsk L total_interference_bound A.

We prove that A is also in the concrete search space.
      Lemma A_is_in_concrete_search_space:
        is_in_search_space A.

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

1 subgoal (ID 1405)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  ============================
  is_in_search_space A

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


      Proof.
        move: H_A_is_in_abstract_search_space ⇒ [INSP | [/andP [POSA LTL] [x [LTx INSP2]]]].

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

2 subgoals (ID 1419)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  INSP : A = 0
  ============================
  is_in_search_space A

subgoal 2 (ID 1489) is:
 is_in_search_space A

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


        - rewrite INSP.

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

2 subgoals (ID 1491)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  INSP : A = 0
  ============================
  is_in_search_space 0

subgoal 2 (ID 1489) is:
 is_in_search_space A

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


          apply/andP; split; first by done.

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

2 subgoals (ID 1518)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  INSP : A = 0
  ============================
  task_rbf 0 != task_rbf (0 + ε)

subgoal 2 (ID 1489) is:
 is_in_search_space A

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


          rewrite neq_ltn; apply/orP; left.

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

2 subgoals (ID 1548)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  INSP : A = 0
  ============================
  task_rbf 0 < task_rbf (0 + ε)

subgoal 2 (ID 1489) is:
 is_in_search_space A

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


          rewrite {1}/task_rbf; erewrite task_rbf_0_zero; eauto 2; try done.

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

2 subgoals (ID 1552)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  INSP : A = 0
  ============================
  0 < task_rbf (0 + ε)

subgoal 2 (ID 1489) is:
 is_in_search_space A

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


          rewrite add0n /task_rbf; apply leq_trans with (task_cost tsk).

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

3 subgoals (ID 1611)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  INSP : A = 0
  ============================
  0 < task_cost tsk

subgoal 2 (ID 1612) is:
 task_cost tsk <= task_request_bound_function tsk ε
subgoal 3 (ID 1489) is:
 is_in_search_space A

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


          + by apply leq_trans with (job_cost j); rewrite -?H_job_of_tsk; eauto 2.

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

2 subgoals (ID 1612)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  INSP : A = 0
  ============================
  task_cost tsk <= task_request_bound_function tsk ε

subgoal 2 (ID 1489) is:
 is_in_search_space A

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


          + eapply task_rbf_1_ge_task_cost; eauto 2.

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

1 subgoal (ID 1489)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  POSA : 0 < A
  LTL : A < L
  x : nat
  LTx : x < L
  INSP2 : total_interference_bound tsk (A - ε) x <>
          total_interference_bound tsk A x
  ============================
  is_in_search_space A

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


        - apply/andP; split; first by done.

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

1 subgoal (ID 1673)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  POSA : 0 < A
  LTL : A < L
  x : nat
  LTx : x < L
  INSP2 : total_interference_bound tsk (A - ε) x <>
          total_interference_bound tsk A x
  ============================
  task_rbf A != task_rbf (A + ε)

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


          apply/negP; intros EQ; move: EQ ⇒ /eqP EQ.

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

1 subgoal (ID 1731)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  POSA : 0 < A
  LTL : A < L
  x : nat
  LTx : x < L
  INSP2 : total_interference_bound tsk (A - ε) x <>
          total_interference_bound tsk A x
  EQ : task_rbf A = task_rbf (A + ε)
  ============================
  False

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


          apply INSP2.

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

1 subgoal (ID 1732)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  POSA : 0 < A
  LTL : A < L
  x : nat
  LTx : x < L
  INSP2 : total_interference_bound tsk (A - ε) x <>
          total_interference_bound tsk A x
  EQ : task_rbf A = task_rbf (A + ε)
  ============================
  total_interference_bound tsk (A - ε) x = total_interference_bound tsk A x

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


          unfold total_interference_bound in ×.

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

1 subgoal (ID 1734)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      (fun (tsk : Task) 
                                         (A : nat) 
                                         (Δ : duration) =>
                                       task_rbf (A + ε) - task_cost tsk +
                                       IBF Δ) A
  POSA : 0 < A
  LTL : A < L
  x : nat
  LTx : x < L
  INSP2 : task_rbf (A - ε + ε) - task_cost tsk + IBF x <>
          task_rbf (A + ε) - task_cost tsk + IBF x
  EQ : task_rbf A = task_rbf (A + ε)
  ============================
  task_rbf (A - ε + ε) - task_cost tsk + IBF x =
  task_rbf (A + ε) - task_cost tsk + IBF x

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


          rewrite subn1 addn1 prednK; last by done.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1744)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      (fun (tsk : Task) 
                                         (A : nat) 
                                         (Δ : duration) =>
                                       task_rbf (A + ε) - task_cost tsk +
                                       IBF Δ) A
  POSA : 0 < A
  LTL : A < L
  x : nat
  LTx : x < L
  INSP2 : task_rbf (A - ε + ε) - task_cost tsk + IBF x <>
          task_rbf (A + ε) - task_cost tsk + IBF x
  EQ : task_rbf A = task_rbf (A + ε)
  ============================
  task_rbf A - task_cost tsk + IBF x =
  task_rbf (A + ε) - task_cost tsk + IBF x

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



            by rewrite -EQ.

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

No more subgoals.

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


      Qed.

Then, there exists a solution for the response-time recurrence (in the abstract sense).
      Corollary correct_search_space:
         (F : duration),
          A + F = task_rbf (A + ε) - (task_cost tsk - task_run_to_completion_threshold tsk) + IBF (A + F)
          F + (task_cost tsk - task_run_to_completion_threshold tsk) R.

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

1 subgoal (ID 1417)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  ============================
  exists F : duration,
    A + F =
    task_rbf (A + ε) - (task_cost tsk - task_run_to_completion_threshold tsk) +
    IBF (A + F) /\
    F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R

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


      Proof.
        move: (H_R_is_maximum A) ⇒ FIX.

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

1 subgoal (ID 1419)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  FIX : is_in_search_space A ->
        exists F : duration,
          A + F =
          priority_inversion_bound +
          (task_rbf (A + ε) -
           (task_cost tsk - task_run_to_completion_threshold tsk)) +
          total_ohep_rbf (A + F) /\
          F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
  ============================
  exists F : duration,
    A + F =
    task_rbf (A + ε) - (task_cost tsk - task_run_to_completion_threshold tsk) +
    IBF (A + F) /\
    F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R

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


        feed FIX; first by apply A_is_in_concrete_search_space.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1425)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  FIX : exists F : duration,
          A + F =
          priority_inversion_bound +
          (task_rbf (A + ε) -
           (task_cost tsk - task_run_to_completion_threshold tsk)) +
          total_ohep_rbf (A + F) /\
          F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
  ============================
  exists F : duration,
    A + F =
    task_rbf (A + ε) - (task_cost tsk - task_run_to_completion_threshold tsk) +
    IBF (A + F) /\
    F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R

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


        move: FIX ⇒ [F [FIX NEQ]].

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

1 subgoal (ID 1448)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  F : duration
  FIX : A + F =
        priority_inversion_bound +
        (task_rbf (A + ε) -
         (task_cost tsk - task_run_to_completion_threshold tsk)) +
        total_ohep_rbf (A + F)
  NEQ : F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
  ============================
  exists F0 : duration,
    A + F0 =
    task_rbf (A + ε) - (task_cost tsk - task_run_to_completion_threshold tsk) +
    IBF (A + F0) /\
    F0 + (task_cost tsk - task_run_to_completion_threshold tsk) <= R

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


         F; split; last by done.

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

1 subgoal (ID 1452)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  F : duration
  FIX : A + F =
        priority_inversion_bound +
        (task_rbf (A + ε) -
         (task_cost tsk - task_run_to_completion_threshold tsk)) +
        total_ohep_rbf (A + F)
  NEQ : F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
  ============================
  A + F =
  task_rbf (A + ε) - (task_cost tsk - task_run_to_completion_threshold tsk) +
  IBF (A + F)

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


        apply/eqP; rewrite {1}FIX.

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

1 subgoal (ID 1510)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_of_tsk : job_task j = tsk
  H_job_cost_positive : job_cost_positive j
  total_interference_bound := fun (tsk : Task) (A : nat) (Δ : duration) =>
                              task_rbf (A + ε) - task_cost tsk + IBF Δ
   : Task -> nat -> duration -> nat
  A : duration
  H_A_is_in_abstract_search_space : search_space.is_in_search_space tsk L
                                      total_interference_bound A
  F : duration
  FIX : A + F =
        priority_inversion_bound +
        (task_rbf (A + ε) -
         (task_cost tsk - task_run_to_completion_threshold tsk)) +
        total_ohep_rbf (A + F)
  NEQ : F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
  ============================
  priority_inversion_bound +
  (task_rbf (A + ε) - (task_cost tsk - task_run_to_completion_threshold tsk)) +
  total_ohep_rbf (A + F) ==
  task_rbf (A + ε) - (task_cost tsk - task_run_to_completion_threshold tsk) +
  IBF (A + F)

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


          by rewrite addnA [_ + priority_inversion_bound]addnC -!addnA.

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

No more subgoals.

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


      Qed.

    End SolutionOfResponseTimeRecurrenceExists.

  End FillingOutHypothesesOfAbstractRTATheorem.

Final Theorem

Based on the properties established above, we apply the abstract analysis framework to infer that [R] is a response-time bound for [tsk].
  Theorem uniprocessor_response_time_bound_fp:
    response_time_bounded_by tsk R.

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

1 subgoal (ID 1366)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  ============================
  response_time_bounded_by tsk R

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


  Proof.
    intros js ARRs TSKs.

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

1 subgoal (ID 1370)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  js : Job
  ARRs : arrives_in arr_seq js
  TSKs : job_task js = tsk
  ============================
  job_response_time_bound sched js R

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


    move: (posnP (@job_cost _ H3 js)) ⇒ [ZERO|POS].

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

2 subgoals (ID 1387)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  js : Job
  ARRs : arrives_in arr_seq js
  TSKs : job_task js = tsk
  ZERO : job_cost js = 0
  ============================
  job_response_time_bound sched js R

subgoal 2 (ID 1388) is:
 job_response_time_bound sched js R

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


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

1 subgoal (ID 1387)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  js : Job
  ARRs : arrives_in arr_seq js
  TSKs : job_task js = tsk
  ZERO : job_cost js = 0
  ============================
  job_response_time_bound sched js R

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


by rewrite /job_response_time_bound /completed_by ZERO.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1388)

subgoal 1 (ID 1388) is:
 job_response_time_bound sched js R

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


}

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

1 subgoal (ID 1388)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  js : Job
  ARRs : arrives_in arr_seq js
  TSKs : job_task js = tsk
  POS : 0 < job_cost js
  ============================
  job_response_time_bound sched js R

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


    eapply uniprocessor_response_time_bound_seq; eauto 3.

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

5 focused subgoals
(shelved: 4) (ID 1431)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  js : Job
  ARRs : arrives_in arr_seq js
  TSKs : job_task js = tsk
  POS : 0 < job_cost js
  ============================
  work_conserving arr_seq sched tsk ?interference ?interfering_workload

subgoal 2 (ID 1433) is:
 interference_and_workload_consistent_with_sequential_tasks arr_seq sched tsk
   ?interference ?interfering_workload
subgoal 3 (ID 1434) is:
 busy_intervals_are_bounded_by arr_seq sched tsk ?interference
   ?interfering_workload ?L
subgoal 4 (ID 1435) is:
 task_interference_is_bounded_by arr_seq sched tsk 
   ?interference ?interfering_workload ?task_interference_bound_function
subgoal 5 (ID 1436) is:
 forall A : duration,
 search_space.is_in_search_space tsk ?L
   (fun (tsk0 : Task) (A0 Δ : duration) =>
    task_request_bound_function tsk (A0 + ε) - task_cost tsk0 +
    ?task_interference_bound_function tsk0 A0 Δ) A ->
 exists F : duration,
   A + F =
   task_request_bound_function tsk (A + ε) -
   (task_cost tsk - task_run_to_completion_threshold tsk) +
   ?task_interference_bound_function tsk A (A + F) /\
   F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R

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


    - by apply instantiated_i_and_w_are_consistent_with_schedule.
(* ----------------------------------[ coqtop ]---------------------------------

4 focused subgoals
(shelved: 2) (ID 1433)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  js : Job
  ARRs : arrives_in arr_seq js
  TSKs : job_task js = tsk
  POS : 0 < job_cost js
  ============================
  interference_and_workload_consistent_with_sequential_tasks arr_seq sched
    tsk interference interfering_workload

subgoal 2 (ID 1434) is:
 busy_intervals_are_bounded_by arr_seq sched tsk interference
   interfering_workload ?L
subgoal 3 (ID 1435) is:
 task_interference_is_bounded_by arr_seq sched tsk interference
   interfering_workload ?task_interference_bound_function
subgoal 4 (ID 1436) is:
 forall A : duration,
 search_space.is_in_search_space tsk ?L
   (fun (tsk0 : Task) (A0 Δ : duration) =>
    task_request_bound_function tsk (A0 + ε) - task_cost tsk0 +
    ?task_interference_bound_function tsk0 A0 Δ) A ->
 exists F : duration,
   A + F =
   task_request_bound_function tsk (A + ε) -
   (task_cost tsk - task_run_to_completion_threshold tsk) +
   ?task_interference_bound_function tsk A (A + F) /\
   F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R

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


    - by apply instantiated_interference_and_workload_consistent_with_sequential_tasks.
(* ----------------------------------[ coqtop ]---------------------------------

3 focused subgoals
(shelved: 2) (ID 1434)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  js : Job
  ARRs : arrives_in arr_seq js
  TSKs : job_task js = tsk
  POS : 0 < job_cost js
  ============================
  busy_intervals_are_bounded_by arr_seq sched tsk interference
    interfering_workload ?L

subgoal 2 (ID 1435) is:
 task_interference_is_bounded_by arr_seq sched tsk interference
   interfering_workload ?task_interference_bound_function
subgoal 3 (ID 1436) is:
 forall A : duration,
 search_space.is_in_search_space tsk ?L
   (fun (tsk0 : Task) (A0 Δ : duration) =>
    task_request_bound_function tsk (A0 + ε) - task_cost tsk0 +
    ?task_interference_bound_function tsk0 A0 Δ) A ->
 exists F : duration,
   A + F =
   task_request_bound_function tsk (A + ε) -
   (task_cost tsk - task_run_to_completion_threshold tsk) +
   ?task_interference_bound_function tsk A (A + F) /\
   F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R

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


    - by apply instantiated_busy_intervals_are_bounded.

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

2 focused subgoals
(shelved: 1) (ID 1435)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  js : Job
  ARRs : arrives_in arr_seq js
  TSKs : job_task js = tsk
  POS : 0 < job_cost js
  ============================
  task_interference_is_bounded_by arr_seq sched tsk interference
    interfering_workload ?task_interference_bound_function

subgoal 2 (ID 1436) is:
 forall A : duration,
 search_space.is_in_search_space tsk L
   (fun (tsk0 : Task) (A0 Δ : duration) =>
    task_request_bound_function tsk (A0 + ε) - task_cost tsk0 +
    ?task_interference_bound_function tsk0 A0 Δ) A ->
 exists F : duration,
   A + F =
   task_request_bound_function tsk (A + ε) -
   (task_cost tsk - task_run_to_completion_threshold tsk) +
   ?task_interference_bound_function tsk A (A + F) /\
   F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R

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


    - by apply instantiated_task_interference_is_bounded.

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

1 subgoal (ID 1436)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable 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
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  work_conserving_ab := work_conserving arr_seq sched
   : Task -> (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  work_conserving_cl := work_conserving.work_conserving arr_seq sched : Prop
  H_work_conserving : work_conserving_cl
  H_sequential_tasks : sequential_tasks sched
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  H_all_jobs_from_taskset : all_jobs_from_taskset 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
  H_valid_preemption_model : valid_preemption_model arr_seq sched
  H_valid_run_to_completion_threshold : valid_task_run_to_completion_threshold
                                          arr_seq tsk
  H6 : FP_policy Task
  H_priority_is_reflexive : reflexive_priorities
  job_pending_at := pending sched : Job -> instant -> bool
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  job_backlogged_at := backlogged sched : Job -> instant -> bool
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  task_rbf := task_request_bound_function tsk : duration -> nat
  total_hep_rbf := total_hep_request_bound_function_FP ts tsk
   : duration -> nat
  total_ohep_rbf := total_ohep_request_bound_function_FP ts tsk
   : duration -> nat
  priority_inversion_bound : duration
  H_priority_inversion_is_bounded : priority_inversion_is_bounded_by arr_seq
                                      sched tsk priority_inversion_bound
  L : duration
  H_L_positive : 0 < L
  H_fixed_point : L = priority_inversion_bound + total_hep_rbf L
  R : duration
  H_R_is_maximum : forall A : duration,
                   is_in_search_space A ->
                   exists F : duration,
                     A + F =
                     priority_inversion_bound +
                     (task_rbf (A + ε) -
                      (task_cost tsk - task_run_to_completion_threshold tsk)) +
                     total_ohep_rbf (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  interference := fun j : Job => [eta ideal_jlfp_rta.interference sched j]
   : Job -> instant -> bool
  interfering_workload := fun j : Job =>
                          [eta ideal_jlfp_rta.interfering_workload arr_seq
                                 sched j] : Job -> instant -> nat
  IBF := fun R : duration => priority_inversion_bound + total_ohep_rbf R
   : duration -> nat
  js : Job
  ARRs : arrives_in arr_seq js
  TSKs : job_task js = tsk
  POS : 0 < job_cost js
  ============================
  forall A : duration,
  search_space.is_in_search_space tsk L
    (fun (tsk0 : Task) (A0 Δ : duration) =>
     task_request_bound_function tsk (A0 + ε) - task_cost tsk0 +
     (fun=> (fun=> [eta IBF])) tsk0 A0 Δ) A ->
  exists F : duration,
    A + F =
    task_request_bound_function tsk (A + ε) -
    (task_cost tsk - task_run_to_completion_threshold tsk) +
    (fun=> (fun=> [eta IBF])) tsk A (A + F) /\
    F + (task_cost tsk - task_run_to_completion_threshold tsk) <= R

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


    - by eapply correct_search_space; eauto 2.
(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

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


  Qed.

End AbstractRTAforFPwithArrivalCurves.