Library prosa.analysis.abstract.abstract_rta


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

Welcome to Coq 8.11.2 (June 2020)

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


Require Export prosa.analysis.definitions.schedulability.
Require Export prosa.analysis.abstract.search_space.
Require Export prosa.analysis.abstract.run_to_completion.

From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq path fintype bigop.

Abstract Response-Time Analysis

In this module, we propose the general framework for response-time analysis (RTA) of uni-processor scheduling of real-time tasks with arbitrary arrival models. We prove that the maximum (with respect to the set of offsets) among the solutions of the response-time bound recurrence is a response time bound for [tsk]. Note that in this section we do not rely on any hypotheses about job sequentiality.
Section Abstract_RTA.

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 kind of uni-service ideal processor state model.
  Context {PState : Type}.
  Context `{ProcessorState Job PState}.
  Hypothesis H_ideal_progress_proc_model : ideal_progress_proc_model PState.
  Hypothesis H_unit_service_proc_model : unit_service_proc_model PState.

Consider any arrival sequence with consistent, non-duplicate arrivals...
Next, consider any schedule of this arrival sequence...
... where jobs do not execute before their arrival nor after completion.
Assume that the job costs are no larger than the task costs.
Consider a task set ts...
  Variable ts : list Task.

... and a task [tsk] of 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.
Let's define some local names for clarity.
Assume we are provided with abstract functions for interference and interfering workload.
We assume that the scheduler is work-conserving.
For simplicity, let's define some local names.
Let L be a constant which bounds any busy interval of task [tsk].
Next, assume that interference_bound_function is a bound on the interference incurred by jobs of task [tsk].
For simplicity, let's define a local name for the search space.
Consider any value [R] 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 tsk - task_run_to_completion_threshold tsk) <= R].
  Variable R: nat.
  Hypothesis H_R_is_maximum:
     A,
      is_in_search_space A
       F,
        A + F = task_run_to_completion_threshold tsk
                + interference_bound_function tsk A (A + F)
        F + (task_cost tsk - task_run_to_completion_threshold tsk) R.

In this section we show a detailed proof of the main theorem that establishes that R is a response-time bound of task [tsk].
  Section ProofOfTheorem.

Consider any job j of [tsk] with positive cost.
    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.

Assume we have a busy interval [t1, t2) of job j that is bounded by L.
    Variable t1 t2: instant.
    Hypothesis H_busy_interval: busy_interval j t1 t2.

Let's define A as a relative arrival time of job j (with respect to time t1).
    Let A := job_arrival j - t1.

In order to prove that R is a response-time bound of job j, we use hypothesis H_R_is_maximum. Note that the relative arrival time (A) is not necessarily from the search space. However, earlier we have proven that for any A there exists another [A_sp] from the search space that shares the same IBF value. Moreover, we've also shown that there exists an [F_sp] such that [F_sp] is a solution of the response time recurrence for parameter [A_sp]. Thus, despite the fact that the relative arrival time may not lie in the search space, we can still use the assumption H_R_is_maximum.
More formally, consider any [A_sp] and [F_sp] such that:..
    Variable A_sp F_sp : duration.

(a) [A_sp] is less than or equal to [A]...
    Hypothesis H_A_gt_Asp : A_sp A.

(b) [interference_bound_function(A, x)] is equal to [interference_bound_function(A_sp, x)] for all [x] less than [L]...
    Hypothesis H_equivalent :
      are_equivalent_at_values_less_than
        (interference_bound_function tsk A)
        (interference_bound_function tsk A_sp) L.

(c) [A_sp] is in the search space, ...
(d) [A_sp + F_sp] is a solution of the response time recurrence...
(e) and finally, [F_sp + (task_last - ε)] is no greater than R.
In this section, we consider the case where the solution is so large that the value of [t1 + A_sp + F_sp] goes beyond the busy interval. Although this case may be impossible in some scenarios, it can be easily proven, since any job that completes by the end of the busy interval remains completed.
    Section FixpointOutsideBusyInterval.

By assumption, suppose that t2 is less than or equal to [t1 + A_sp + F_sp].
      Hypothesis H_big_fixpoint_solution : t2 t1 + (A_sp + F_sp).

Then we prove that [job_arrival j + R] is no less than [t2].
      Lemma t2_le_arrival_plus_R:
        t2 job_arrival j + R.

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

1 subgoal (ID 702)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_big_fixpoint_solution : t2 <= t1 + (A_sp + F_sp)
  ============================
  t2 <= job_arrival j + R

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


      Proof.
        move: H_busy_interval ⇒ [[/andP [GT LT] [QT1 NTQ]] QT2].

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

1 subgoal (ID 773)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_big_fixpoint_solution : t2 <= t1 + (A_sp + F_sp)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  ============================
  t2 <= job_arrival j + R

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


        apply leq_trans with (t1 + (A_sp + F_sp)); first by done.

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

1 subgoal (ID 775)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_big_fixpoint_solution : t2 <= t1 + (A_sp + F_sp)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  ============================
  t1 + (A_sp + F_sp) <= job_arrival j + R

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


        apply leq_trans with (t1 + A + F_sp).

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

2 subgoals (ID 776)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_big_fixpoint_solution : t2 <= t1 + (A_sp + F_sp)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  ============================
  t1 + (A_sp + F_sp) <= t1 + A + F_sp

subgoal 2 (ID 777) is:
 t1 + A + F_sp <= job_arrival j + R

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


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

1 subgoal (ID 776)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_big_fixpoint_solution : t2 <= t1 + (A_sp + F_sp)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  ============================
  t1 + (A_sp + F_sp) <= t1 + A + F_sp

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


by rewrite !addnA leq_add2r leq_add2l.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 777)

subgoal 1 (ID 777) is:
 t1 + A + F_sp <= job_arrival j + R

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


}

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

1 subgoal (ID 777)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_big_fixpoint_solution : t2 <= t1 + (A_sp + F_sp)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  ============================
  t1 + A + F_sp <= job_arrival j + R

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


        rewrite /A subnKC; last by done.

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

1 subgoal (ID 805)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_big_fixpoint_solution : t2 <= t1 + (A_sp + F_sp)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  ============================
  job_arrival j + F_sp <= job_arrival j + R

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


        rewrite leq_add2l.

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

1 subgoal (ID 811)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_big_fixpoint_solution : t2 <= t1 + (A_sp + F_sp)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  ============================
  F_sp <= R

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


          by apply leq_trans with (F_sp + (task_cost tsk - task_run_to_completion_threshold tsk));
            first rewrite leq_addr.

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

No more subgoals.

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


      Qed.

But since we know that the job is completed by the end of its busy interval, we can show that job j is completed by [job arrival j + R].
      Lemma job_completed_by_arrival_plus_R_1:
        completed_by sched j (job_arrival j + R).

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

1 subgoal (ID 709)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_big_fixpoint_solution : t2 <= t1 + (A_sp + F_sp)
  ============================
  completed_by sched j (job_arrival j + R)

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


      Proof.
        move: H_busy_interval ⇒ [[/andP [GT LT] [QT1 NTQ]] QT2].

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

1 subgoal (ID 780)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_big_fixpoint_solution : t2 <= t1 + (A_sp + F_sp)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  ============================
  completed_by sched j (job_arrival j + R)

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


        apply completion_monotonic with t2; try done.

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

2 subgoals (ID 785)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_big_fixpoint_solution : t2 <= t1 + (A_sp + F_sp)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  ============================
  t2 <= job_arrival j + R

subgoal 2 (ID 786) is:
 completed_by sched j t2

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


        apply t2_le_arrival_plus_R.

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

1 subgoal (ID 786)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_big_fixpoint_solution : t2 <= t1 + (A_sp + F_sp)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  ============================
  completed_by sched j t2

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


        eapply job_completes_within_busy_interval; eauto 2.

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

No more subgoals.

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


      Qed.

    End FixpointOutsideBusyInterval.

In this section, we consider the complementary case where [t1 + A_sp + F_sp] lies inside the busy interval.
    Section FixpointInsideBusyInterval.

So, assume that [t1 + A_sp + F_sp] is less than t2.
      Hypothesis H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2.

Next, let's consider two other cases: CASE 1: the value of the fix-point is no less than the relative arrival time of job [j].
      Section FixpointIsNoLessThanArrival.

Suppose that [A_sp + F_sp] is no less than relative arrival of job [j].
In this section, we prove that the fact that job [j] is not completed by time [job_arrival j + R] leads to a contradiction. Which in turn implies that the opposite is true -- job [j] completes by time [job_arrival j + R].
        Section ProofByContradiction.

Recall that by lemma "solution_for_A_exists" there is a solution [F] of the response-time recurrence for the given relative arrival time [A] (which is not necessarily from the search space).
Thus, consider a constant [F] such that:..
          Variable F : duration.
(a) the sum of [A_sp] and [F_sp] is equal to the sum of [A] and [F]...
          Hypothesis H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F.
(b) [F] is at mo1st [F_sp]...
          Hypothesis H_F_le_Fsp : F F_sp.
(c) and [A + F] is a solution for the response-time recurrence for [A].
          Hypothesis H_A_F_fixpoint:
            A + F = task_run_to_completion_threshold tsk + interference_bound_function tsk A (A + F).

Next, we assume that job [j] is not completed by time [job_arrival j + R].
          Hypothesis H_j_not_completed : ~~ completed_by sched j (job_arrival j + R).

Some additional reasoning is required since the term [task_cost tsk - task_run_to_completion_threshold tsk] does not necessarily bound the term [job_cost j - job_run_to_completion_threshold j]. That is, a job can have a small run-to-completion threshold, thereby becoming non-preemptive much earlier than guaranteed according to task run-to-completion threshold, while simultaneously executing the last non-preemptive segment that is longer than [task_cost tsk - task_run_to_completion_threshold tsk] (e.g., this is possible in the case of floating non-preemptive sections).
In this case we cannot directly apply lemma "j_receives_at_least_run_to_completion_threshold". Therefore we introduce two temporal notions of the last non-preemptive region of job j and an execution optimism. We use these notions inside this proof, so we define them only locally.
Let the last non-preemptive region of job [j] (last) be the difference between the cost of the job and the [j]'s run-to-completion threshold (i.e. [job_cost j - job_run_to_completion_threshold j]). We know that after j has reached its run-to-completion threshold, it will additionally be executed [job_last j] units of time.
          Let job_last := job_cost j - job_run_to_completion_threshold j.

And let execution optimism (optimism) be the difference between the [tsk]'s run-to-completion threshold and the [j]'s run-to-completion threshold (i.e. [task_run_to_completion_threshold - job_run_to_completion_threshold]). Intuitively, optimism is how much earlier job j has received its run-to-completion threshold than it could at worst.
From lemma "j_receives_at_least_run_to_completion_threshold" with parameters [progress_of_job := job_run_to_completion_threshold j] and [delta := (A + F) - optimism)] we know that service of [j] by time [t1 + (A + F) - optimism] is no less than [job_run_to_completion_threshold j]. Hence, job [j] is completed by time [t1 + (A + F) - optimism + last].
          Lemma j_is_completed_by_t1_A_F_optimist_last :
            completed_by sched j (t1 + (A + F - optimism) + job_last).

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

1 subgoal (ID 726)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  ============================
  completed_by sched j (t1 + (A + F - optimism) + job_last)

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


          Proof.
            have HelpAuto: m n, n n + m; first by intros; rewrite leq_addr.

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

1 subgoal (ID 731)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  ============================
  completed_by sched j (t1 + (A + F - optimism) + job_last)

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


            move: H_busy_interval ⇒ [[/andP [GT LT] _] _].

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

1 subgoal (ID 798)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  completed_by sched j (t1 + (A + F - optimism) + job_last)

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


            have ESERV :=
              @j_receives_at_least_run_to_completion_threshold
                _ _ H1 H2 H3 PState H5 _ _ arr_seq sched tsk interference interfering_workload
                _ j _ _ _ t1 t2 _ (job_run_to_completion_threshold j) _ ((A + F) - optimism).
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 826)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : ideal_progress_proc_model PState ->
          unit_service_proc_model PState ->
          definitions.work_conserving arr_seq sched tsk interference
            interfering_workload ->
          arrives_in arr_seq j ->
          job_task j = tsk ->
          job_cost_positive j ->
          definitions.busy_interval sched interference interfering_workload j
            t1 t2 ->
          forall (j0 : JobCost Job) (j1 : JobPreemptable Job),
          job_run_to_completion_threshold j <= job_cost j ->
          job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  completed_by sched j (t1 + (A + F - optimism) + job_last)

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


            feed_n 7 ESERV; eauto 2.

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

1 subgoal (ID 868)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : forall (j0 : JobCost Job) (j1 : JobPreemptable Job),
          job_run_to_completion_threshold j <= job_cost j ->
          job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  completed_by sched j (t1 + (A + F - optimism) + job_last)

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


            specialize (ESERV H3 H4).

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

1 subgoal (ID 882)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j <= job_cost j ->
          job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  completed_by sched j (t1 + (A + F - optimism) + job_last)

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


            feed_n 2 ESERV.

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

3 subgoals (ID 883)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j <= job_cost j ->
          job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  job_run_to_completion_threshold j <= job_cost j

subgoal 2 (ID 889) is:
 job_run_to_completion_threshold j +
 definitions.cumul_interference interference j t1 (t1 + (A + F - optimism)) <=
 A + F - optimism
subgoal 3 (ID 894) is:
 completed_by sched j (t1 + (A + F - optimism) + job_last)

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


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

1 subgoal (ID 883)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j <= job_cost j ->
          job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  job_run_to_completion_threshold j <= job_cost j

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


eapply job_run_to_completion_threshold_le_job_cost; eauto.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 889)

subgoal 1 (ID 889) is:
 job_run_to_completion_threshold j +
 definitions.cumul_interference interference j t1 (t1 + (A + F - optimism)) <=
 A + F - optimism
subgoal 2 (ID 894) is:
 completed_by sched j (t1 + (A + F - optimism) + job_last)

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


}

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

2 subgoals (ID 889)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  job_run_to_completion_threshold j +
  definitions.cumul_interference interference j t1 (t1 + (A + F - optimism)) <=
  A + F - optimism

subgoal 2 (ID 894) is:
 completed_by sched j (t1 + (A + F - optimism) + job_last)

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


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

1 subgoal (ID 889)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  job_run_to_completion_threshold j +
  definitions.cumul_interference interference j t1 (t1 + (A + F - optimism)) <=
  A + F - optimism

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


rewrite {2}H_A_F_fixpoint.

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

1 subgoal (ID 899)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  job_run_to_completion_threshold j +
  definitions.cumul_interference interference j t1 (t1 + (A + F - optimism)) <=
  task_run_to_completion_threshold tsk +
  interference_bound_function tsk A (A + F) - optimism

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


              rewrite /definitions.cumul_interference.

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

1 subgoal (ID 902)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  job_run_to_completion_threshold j +
  \sum_(t1 <= t < t1 + (A + F - optimism)) interference j t <=
  task_run_to_completion_threshold tsk +
  interference_bound_function tsk A (A + F) - optimism

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


              rewrite -[in X in _ X]subh1; last by rewrite leq_subr.

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

1 subgoal (ID 913)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  job_run_to_completion_threshold j +
  \sum_(t1 <= t < t1 + (A + F - optimism)) interference j t <=
  task_run_to_completion_threshold tsk - optimism +
  interference_bound_function tsk A (A + F)

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


              rewrite {2}/optimism.

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

1 subgoal (ID 919)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  job_run_to_completion_threshold j +
  \sum_(t1 <= t < t1 + (A + F - optimism)) interference j t <=
  task_run_to_completion_threshold tsk -
  (task_run_to_completion_threshold tsk - job_run_to_completion_threshold j) +
  interference_bound_function tsk A (A + F)

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


              rewrite subKn; last by apply H_valid_run_to_completion_threshold.

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

1 subgoal (ID 924)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  job_run_to_completion_threshold j +
  \sum_(t1 <= t < t1 + (A + F - optimism)) interference j t <=
  job_run_to_completion_threshold j +
  interference_bound_function tsk A (A + F)

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


              rewrite leq_add2l.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 936)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  \sum_(t1 <= t < t1 + (A + F - optimism)) interference j t <=
  interference_bound_function tsk A (A + F)

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



              apply leq_trans with (cumul_interference j t1 (t1 + (A + F))).

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

2 subgoals (ID 937)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  \sum_(t1 <= t < t1 + (A + F - optimism)) interference j t <=
  cumul_interference j t1 (t1 + (A + F))

subgoal 2 (ID 938) is:
 cumul_interference j t1 (t1 + (A + F)) <=
 interference_bound_function tsk A (A + F)

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


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

1 subgoal (ID 937)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  \sum_(t1 <= t < t1 + (A + F - optimism)) interference j t <=
  cumul_interference j t1 (t1 + (A + F))

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


rewrite /cumul_interference /definitions.cumul_interference;
                  rewrite [in X in _ X](@big_cat_nat _ _ _ (t1 + (A + F - optimism))) //=.

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

1 subgoal (ID 972)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  t1 + (A + F - optimism) <= t1 + (A + F)

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


                  by rewrite leq_add2l leq_subr.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 938)

subgoal 1 (ID 938) is:
 cumul_interference j t1 (t1 + (A + F)) <=
 interference_bound_function tsk A (A + F)
subgoal 2 (ID 894) is:
 completed_by sched j (t1 + (A + F - optimism) + job_last)

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


}

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

1 subgoal (ID 938)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  cumul_interference j t1 (t1 + (A + F)) <=
  interference_bound_function tsk A (A + F)

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


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

1 subgoal (ID 938)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  cumul_interference j t1 (t1 + (A + F)) <=
  interference_bound_function tsk A (A + F)

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


apply H_job_interference_is_bounded with t2; try done.

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

2 subgoals (ID 1007)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  t1 + (A + F) < t2

subgoal 2 (ID 1008) is:
 ~~ completed_by sched j (t1 + (A + F))

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


                - by rewrite -H_Asp_Fsp_eq_A_F.

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

1 subgoal (ID 1008)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  ~~ completed_by sched j (t1 + (A + F))

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


                - apply/negP; intros CONTR.

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

1 subgoal (ID 1076)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  CONTR : completed_by sched j (t1 + (A + F))
  ============================
  False

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


                  move: H_j_not_completed ⇒ /negP C; apply: C.

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

1 subgoal (ID 1109)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  CONTR : completed_by sched j (t1 + (A + F))
  ============================
  completed_by sched j (job_arrival j + R)

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


                  apply completion_monotonic with (t1 + (A + F)); try done.

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

1 subgoal (ID 1114)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  CONTR : completed_by sched j (t1 + (A + F))
  ============================
  t1 + (A + F) <= job_arrival j + R

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


                  rewrite addnA subnKC // leq_add2l.

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

1 subgoal (ID 1177)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  CONTR : completed_by sched j (t1 + (A + F))
  ============================
  F <= R

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


                  apply leq_trans with F_sp; first by done.

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

1 subgoal (ID 1179)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A + F - optimism)) <= A + F - optimism ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  CONTR : completed_by sched j (t1 + (A + F))
  ============================
  F_sp <= R

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


                    by apply leq_trans with (F_sp + (task_cost tsk - task_run_to_completion_threshold tsk)).

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

1 subgoal (ID 894)

subgoal 1 (ID 894) is:
 completed_by sched j (t1 + (A + F - optimism) + job_last)

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


              }

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

1 subgoal (ID 894)

subgoal 1 (ID 894) is:
 completed_by sched j (t1 + (A + F - optimism) + job_last)

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


            }

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

1 subgoal (ID 894)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j <=
          service sched j (t1 + (A + F - optimism))
  ============================
  completed_by sched j (t1 + (A + F - optimism) + job_last)

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


            eapply job_completes_after_reaching_run_to_completion_threshold with (arr_seq0 := arr_seq); eauto 2.

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

No more subgoals.

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


          Qed.

However, [t1 + (A + F) - optimism + last ≤ job_arrival j + R]! To prove this fact we need a few auxiliary inequalities that are needed because we use the truncated subtraction in our development. So, for example [a + (b - c) = a + b - c] only if [b ≥ c].
          Section AuxiliaryInequalities.

Recall that we consider a busy interval of a job [j], and [j] has arrived [A] time units after the beginning the busy interval. From basic properties of a busy interval it follows that job [j] incurs interference at any time instant t ∈ [t1, t1 + A). Therefore [interference_bound_function(tsk, A, A + F)] is at least [A].
            Lemma relative_arrival_le_interference_bound:
              A interference_bound_function tsk A (A + F).

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

1 subgoal (ID 727)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  ============================
  A <= interference_bound_function tsk A (A + F)

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


            Proof.
              have HelpAuto: m n, n n + m; first by intros; rewrite leq_addr.

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

1 subgoal (ID 732)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  ============================
  A <= interference_bound_function tsk A (A + F)

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


              move: H_j_not_completed; clear H_j_not_completed; move ⇒ /negP CONTRc.

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

1 subgoal (ID 766)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  ============================
  A <= interference_bound_function tsk A (A + F)

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


              move: (H_busy_interval) ⇒ [[/andP [GT LT] _] _].

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

1 subgoal (ID 827)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  A <= interference_bound_function tsk A (A + F)

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


              apply leq_trans with (cumul_interference j t1 (t1 + (A+F))).

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

2 subgoals (ID 828)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  A <= cumul_interference j t1 (t1 + (A + F))

subgoal 2 (ID 829) is:
 cumul_interference j t1 (t1 + (A + F)) <=
 interference_bound_function tsk A (A + F)

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


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

1 subgoal (ID 828)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  A <= cumul_interference j t1 (t1 + (A + F))

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


unfold cumul_interference.

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

1 subgoal (ID 831)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  A <= definitions.cumul_interference interference j t1 (t1 + (A + F))

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


                apply leq_trans with
                    (\sum_(t1 t < t1 + A) interference j t); last first.

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

2 subgoals (ID 837)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  \sum_(t1 <= t < t1 + A) interference j t <=
  definitions.cumul_interference interference j t1 (t1 + (A + F))

subgoal 2 (ID 836) is:
 A <= \sum_(t1 <= t < t1 + A) interference j t

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


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

1 subgoal (ID 837)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  \sum_(t1 <= t < t1 + A) interference j t <=
  definitions.cumul_interference interference j t1 (t1 + (A + F))

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


unfold definitions.cumul_interference.

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

1 subgoal (ID 839)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  \sum_(t1 <= t < t1 + A) interference j t <=
  \sum_(t1 <= t < t1 + (A + F)) interference j t

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


                  rewrite [in X in _ X](@big_cat_nat _ _ _ (t1 + A)) //=; last by rewrite addnA.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 836)

subgoal 1 (ID 836) is:
 A <= \sum_(t1 <= t < t1 + A) interference j t
subgoal 2 (ID 829) is:
 cumul_interference j t1 (t1 + (A + F)) <=
 interference_bound_function tsk A (A + F)

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


}

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

1 subgoal (ID 836)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  A <= \sum_(t1 <= t < t1 + A) interference j t

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


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

1 subgoal (ID 836)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  A <= \sum_(t1 <= t < t1 + A) interference j t

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


rewrite -{1}[A](sum_of_ones t1).

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

1 subgoal (ID 902)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  \sum_(t1 <= x < t1 + A) 1 <= \sum_(t1 <= t < t1 + A) interference j t

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


                  rewrite [in X in X _]big_nat_cond [in X in _ X]big_nat_cond.

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

1 subgoal (ID 930)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  \sum_(t1 <= i < t1 + A | (t1 <= i < t1 + A) && true) 1 <=
  \sum_(t1 <= i < t1 + A | (t1 <= i < t1 + A) && true) interference j i

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


                  rewrite leq_sum //.

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

1 subgoal (ID 939)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  forall i : nat, (t1 <= i < t1 + A) && true -> 0 < interference j i

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


                  movet /andP [/andP [NEQ1 NEQ2] _].

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

1 subgoal (ID 1045)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  t : nat
  NEQ1 : t1 <= t
  NEQ2 : t < t1 + A
  ============================
  0 < interference j t

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


                  rewrite lt0b.

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

1 subgoal (ID 1048)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  t : nat
  NEQ1 : t1 <= t
  NEQ2 : t < t1 + A
  ============================
  interference j t

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


                  unfold work_conserving in H_work_conserving.

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

1 subgoal (ID 1049)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : definitions.work_conserving arr_seq sched tsk
                        interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  t : nat
  NEQ1 : t1 <= t
  NEQ2 : t < t1 + A
  ============================
  interference j t

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


                  move: (H_work_conserving j t1 t2 t) ⇒ CONS.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1051)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : definitions.work_conserving arr_seq sched tsk
                        interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  t : nat
  NEQ1 : t1 <= t
  NEQ2 : t < t1 + A
  CONS : arrives_in arr_seq j ->
         job_task j = tsk ->
         0 < job_cost j ->
         definitions.busy_interval sched interference interfering_workload j
           t1 t2 ->
         t1 <= t < t2 -> ~ interference j t <-> scheduled_at sched j t
  ============================
  interference j t

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


                  feed_n 5 CONS; try done.

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

2 subgoals (ID 1076)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : definitions.work_conserving arr_seq sched tsk
                        interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  t : nat
  NEQ1 : t1 <= t
  NEQ2 : t < t1 + A
  CONS : t1 <= t < t2 -> ~ interference j t <-> scheduled_at sched j t
  ============================
  t1 <= t < t2

subgoal 2 (ID 1081) is:
 interference j t

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


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

1 subgoal (ID 1076)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : definitions.work_conserving arr_seq sched tsk
                        interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  t : nat
  NEQ1 : t1 <= t
  NEQ2 : t < t1 + A
  CONS : t1 <= t < t2 -> ~ interference j t <-> scheduled_at sched j t
  ============================
  t1 <= t < t2

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


apply/andP; split; first by done.

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

1 subgoal (ID 1154)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : definitions.work_conserving arr_seq sched tsk
                        interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  t : nat
  NEQ1 : t1 <= t
  NEQ2 : t < t1 + A
  CONS : t1 <= t < t2 -> ~ interference j t <-> scheduled_at sched j t
  ============================
  t < t2

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


                    apply leq_trans with (t1 + A); first by done.

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

1 subgoal (ID 1156)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : definitions.work_conserving arr_seq sched tsk
                        interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  t : nat
  NEQ1 : t1 <= t
  NEQ2 : t < t1 + A
  CONS : t1 <= t < t2 -> ~ interference j t <-> scheduled_at sched j t
  ============================
  t1 + A <= t2

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


                      by rewrite /A subnKC // ltnW.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1081)

subgoal 1 (ID 1081) is:
 interference j t
subgoal 2 (ID 829) is:
 cumul_interference j t1 (t1 + (A + F)) <=
 interference_bound_function tsk A (A + F)

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


                  }

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

1 subgoal (ID 1081)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : definitions.work_conserving arr_seq sched tsk
                        interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  t : nat
  NEQ1 : t1 <= t
  NEQ2 : t < t1 + A
  CONS : ~ interference j t <-> scheduled_at sched j t
  ============================
  interference j t

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


                  move: CONS ⇒ [CONS1 _].

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

1 subgoal (ID 1206)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : definitions.work_conserving arr_seq sched tsk
                        interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  t : nat
  NEQ1 : t1 <= t
  NEQ2 : t < t1 + A
  CONS1 : ~ interference j t -> scheduled_at sched j t
  ============================
  interference j t

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


                  apply/negP; intros CONTR.

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

1 subgoal (ID 1244)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : definitions.work_conserving arr_seq sched tsk
                        interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  t : nat
  NEQ1 : t1 <= t
  NEQ2 : t < t1 + A
  CONS1 : ~ interference j t -> scheduled_at sched j t
  CONTR : ~ interference j t
  ============================
  False

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


                  move: (CONS1 CONTR) ⇒ SCHED; clear CONS1 CONTR.

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

1 subgoal (ID 1247)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : definitions.work_conserving arr_seq sched tsk
                        interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  t : nat
  NEQ1 : t1 <= t
  NEQ2 : t < t1 + A
  SCHED : scheduled_at sched j t
  ============================
  False

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


                  apply H_jobs_must_arrive_to_execute in SCHED.

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

1 subgoal (ID 1248)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : definitions.work_conserving arr_seq sched tsk
                        interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  t : nat
  NEQ1 : t1 <= t
  NEQ2 : t < t1 + A
  SCHED : has_arrived j t
  ============================
  False

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


                  move: NEQ2; rewrite ltnNge; move ⇒ /negP NEQ2; apply: NEQ2.

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

1 subgoal (ID 1287)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : definitions.work_conserving arr_seq sched tsk
                        interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  t : nat
  NEQ1 : t1 <= t
  SCHED : has_arrived j t
  ============================
  t1 + A <= t

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


                    by rewrite subnKC.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 829)

subgoal 1 (ID 829) is:
 cumul_interference j t1 (t1 + (A + F)) <=
 interference_bound_function tsk A (A + F)

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



                }

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

1 subgoal (ID 829)

subgoal 1 (ID 829) is:
 cumul_interference j t1 (t1 + (A + F)) <=
 interference_bound_function tsk A (A + F)

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


              }

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

1 subgoal (ID 829)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  cumul_interference j t1 (t1 + (A + F)) <=
  interference_bound_function tsk A (A + F)

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


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

1 subgoal (ID 829)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  cumul_interference j t1 (t1 + (A + F)) <=
  interference_bound_function tsk A (A + F)

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


apply H_job_interference_is_bounded with t2; try done.

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

2 subgoals (ID 1297)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  t1 + (A + F) < t2

subgoal 2 (ID 1298) is:
 ~~ completed_by sched j (t1 + (A + F))

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


                - by rewrite -H_Asp_Fsp_eq_A_F.

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

1 subgoal (ID 1298)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  CONTRc : ~ completed_by sched j (job_arrival j + R)
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  ~~ completed_by sched j (t1 + (A + F))

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


                - apply/negP; intros CONTR; apply: CONTRc.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1374)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  CONTR : completed_by sched j (t1 + (A + F))
  ============================
  completed_by sched j (job_arrival j + R)

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


                  apply completion_monotonic with (t1 + (A + F)); last by done.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1379)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  CONTR : completed_by sched j (t1 + (A + F))
  ============================
  t1 + (A + F) <= job_arrival j + R

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


                  rewrite !addnA subnKC // leq_add2l.

(* ----------------------------------[ 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
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  CONTR : completed_by sched j (t1 + (A + F))
  ============================
  F <= R

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


                  apply leq_trans with F_sp; first by done.

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

1 subgoal (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
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  CONTR : completed_by sched j (t1 + (A + F))
  ============================
  F_sp <= R

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


                    by apply leq_trans with (F_sp + (task_cost tsk - task_run_to_completion_threshold tsk)).

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

No more subgoals.

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


              }

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

No more subgoals.

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


            Qed.

As two trivial corollaries, we show that [tsk]'s run-to-completion threshold is at most [F_sp]...
            Corollary tsk_run_to_completion_threshold_le_Fsp :
              task_run_to_completion_threshold tsk F_sp.

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

1 subgoal (ID 730)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  ============================
  task_run_to_completion_threshold tsk <= F_sp

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


            Proof.
              have HH : task_run_to_completion_threshold tsk F.

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

2 subgoals (ID 733)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  ============================
  task_run_to_completion_threshold tsk <= F

subgoal 2 (ID 735) is:
 task_run_to_completion_threshold tsk <= F_sp

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


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

1 subgoal (ID 733)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  ============================
  task_run_to_completion_threshold tsk <= F

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


move: H_A_F_fixpointEQ.

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

1 subgoal (ID 737)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  EQ : A + F =
       task_run_to_completion_threshold tsk +
       interference_bound_function tsk A (A + F)
  ============================
  task_run_to_completion_threshold tsk <= F

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


                have L1 := relative_arrival_le_interference_bound.

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

1 subgoal (ID 742)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  EQ : A + F =
       task_run_to_completion_threshold tsk +
       interference_bound_function tsk A (A + F)
  L1 : A <= interference_bound_function tsk A (A + F)
  ============================
  task_run_to_completion_threshold tsk <= F

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


                ssrlia.

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

1 subgoal (ID 735)

subgoal 1 (ID 735) is:
 task_run_to_completion_threshold tsk <= F_sp

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


              }

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

1 subgoal (ID 735)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HH : task_run_to_completion_threshold tsk <= F
  ============================
  task_run_to_completion_threshold tsk <= F_sp

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


              apply leq_trans with F; auto.

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

No more subgoals.

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


            Qed.

... and optimism is at most [F].
            Corollary optimism_le_F :
              optimism F.

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

1 subgoal (ID 731)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  ============================
  optimism <= F

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


            Proof.
              have HH : task_run_to_completion_threshold tsk F.

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

2 subgoals (ID 734)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  ============================
  task_run_to_completion_threshold tsk <= F

subgoal 2 (ID 736) is:
 optimism <= F

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


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

1 subgoal (ID 734)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  ============================
  task_run_to_completion_threshold tsk <= F

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


move: H_A_F_fixpointEQ.

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

1 subgoal (ID 738)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  EQ : A + F =
       task_run_to_completion_threshold tsk +
       interference_bound_function tsk A (A + F)
  ============================
  task_run_to_completion_threshold tsk <= F

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


                have L1 := relative_arrival_le_interference_bound.

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

1 subgoal (ID 743)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  EQ : A + F =
       task_run_to_completion_threshold tsk +
       interference_bound_function tsk A (A + F)
  L1 : A <= interference_bound_function tsk A (A + F)
  ============================
  task_run_to_completion_threshold tsk <= F

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


                ssrlia.

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

1 subgoal (ID 736)

subgoal 1 (ID 736) is:
 optimism <= F

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


              }

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

1 subgoal (ID 736)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  HH : task_run_to_completion_threshold tsk <= F
  ============================
  optimism <= F

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


                by apply leq_trans with (task_run_to_completion_threshold tsk); first rewrite /optimism leq_subr.

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

No more subgoals.

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


            Qed.

          End AuxiliaryInequalities.

Next we show that [t1 + (A + F) - optimism + last] is at most [job_arrival j + R], which is easy to see from the following sequence of inequalities:
[t1 + (A + F) - optimism + last] [≤ job_arrival j + (F - optimism) + job_last] [≤ job_arrival j + (F_sp - optimism) + job_last] [≤ job_arrival j + F_sp + (job_last - optimism)] [≤ job_arrival j + F_sp + job_cost j - task_run_to_completion_threshold tsk] [≤ job_arrival j + F_sp + task_cost tsk - task_run_to_completion_threshold tsk] [≤ job_arrival j + R].
          Lemma t1_A_F_optimist_last_le_arrival_R :
            t1 + (A + F - optimism) + job_last job_arrival j + R.

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

1 subgoal (ID 729)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  ============================
  t1 + (A + F - optimism) + job_last <= job_arrival j + R

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


          Proof.
            move: (H_busy_interval) ⇒ [[/andP [GT LT] _] _].

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

1 subgoal (ID 790)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  t1 + (A + F - optimism) + job_last <= job_arrival j + R

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


            have L1 := tsk_run_to_completion_threshold_le_Fsp.

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

1 subgoal (ID 795)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  ============================
  t1 + (A + F - optimism) + job_last <= job_arrival j + R

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


            have L2 := optimism_le_F.

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

1 subgoal (ID 800)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  ============================
  t1 + (A + F - optimism) + job_last <= job_arrival j + R

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


            apply leq_trans with (job_arrival j + (F - optimism) + job_last).

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

2 subgoals (ID 803)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  ============================
  t1 + (A + F - optimism) + job_last <=
  job_arrival j + (F - optimism) + job_last

subgoal 2 (ID 804) is:
 job_arrival j + (F - optimism) + job_last <= job_arrival j + R

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


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

1 subgoal (ID 803)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  ============================
  t1 + (A + F - optimism) + job_last <=
  job_arrival j + (F - optimism) + job_last

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


rewrite leq_add2r addnBA.

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

2 subgoals (ID 815)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  ============================
  t1 + (A + F) - optimism <= job_arrival j + (F - optimism)

subgoal 2 (ID 816) is:
 optimism <= A + F

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


              - by rewrite /A !addnA subnKC // addnBA.

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

1 subgoal (ID 816)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  ============================
  optimism <= A + F

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


              - by apply leq_trans with F; last rewrite leq_addl.

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

1 subgoal (ID 804)

subgoal 1 (ID 804) is:
 job_arrival j + (F - optimism) + job_last <= job_arrival j + R

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


            }

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

1 subgoal (ID 804)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  ============================
  job_arrival j + (F - optimism) + job_last <= job_arrival j + R

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


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

1 subgoal (ID 804)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  ============================
  job_arrival j + (F - optimism) + job_last <= job_arrival j + R

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


move: H_valid_run_to_completion_threshold ⇒ [PRT1 PRT2].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 882)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  PRT1 : task_rtc_bounded_by_cost tsk
  PRT2 : job_respects_task_rtc arr_seq tsk
  ============================
  job_arrival j + (F - optimism) + job_last <= job_arrival j + R

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


              rewrite -addnA leq_add2l.

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

1 subgoal (ID 894)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  PRT1 : task_rtc_bounded_by_cost tsk
  PRT2 : job_respects_task_rtc arr_seq tsk
  ============================
  F - optimism + job_last <= R

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


              apply leq_trans with (F_sp - optimism + job_last ); first by rewrite leq_add2r leq_sub2r.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 896)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  PRT1 : task_rtc_bounded_by_cost tsk
  PRT2 : job_respects_task_rtc arr_seq tsk
  ============================
  F_sp - optimism + job_last <= R

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


              apply leq_trans with (F_sp + (task_cost tsk - task_run_to_completion_threshold tsk)); last by done.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 913)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  PRT1 : task_rtc_bounded_by_cost tsk
  PRT2 : job_respects_task_rtc arr_seq tsk
  ============================
  F_sp - optimism + job_last <=
  F_sp + (task_cost tsk - task_run_to_completion_threshold tsk)

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


              rewrite /optimism subnBA; last by apply PRT2.

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

1 subgoal (ID 921)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  PRT1 : task_rtc_bounded_by_cost tsk
  PRT2 : job_respects_task_rtc arr_seq tsk
  ============================
  F_sp + job_run_to_completion_threshold j -
  task_run_to_completion_threshold tsk + job_last <=
  F_sp + (task_cost tsk - task_run_to_completion_threshold tsk)

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


              rewrite -subh1 //.

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

1 subgoal (ID 930)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  PRT1 : task_rtc_bounded_by_cost tsk
  PRT2 : job_respects_task_rtc arr_seq tsk
  ============================
  F_sp - task_run_to_completion_threshold tsk +
  job_run_to_completion_threshold j + job_last <=
  F_sp + (task_cost tsk - task_run_to_completion_threshold tsk)

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


              rewrite /job_last.

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

1 subgoal (ID 954)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  PRT1 : task_rtc_bounded_by_cost tsk
  PRT2 : job_respects_task_rtc arr_seq tsk
  ============================
  F_sp - task_run_to_completion_threshold tsk +
  job_run_to_completion_threshold j +
  (job_cost j - job_run_to_completion_threshold j) <=
  F_sp + (task_cost tsk - task_run_to_completion_threshold tsk)

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


              rewrite addnBA; last by eapply job_run_to_completion_threshold_le_job_cost; eauto.

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

1 subgoal (ID 960)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  PRT1 : task_rtc_bounded_by_cost tsk
  PRT2 : job_respects_task_rtc arr_seq tsk
  ============================
  F_sp - task_run_to_completion_threshold tsk +
  job_run_to_completion_threshold j + job_cost j -
  job_run_to_completion_threshold j <=
  F_sp + (task_cost tsk - task_run_to_completion_threshold tsk)

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


              rewrite -subh1; last by rewrite leq_addl.

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

1 subgoal (ID 970)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  PRT1 : task_rtc_bounded_by_cost tsk
  PRT2 : job_respects_task_rtc arr_seq tsk
  ============================
  F_sp - task_run_to_completion_threshold tsk +
  job_run_to_completion_threshold j - job_run_to_completion_threshold j +
  job_cost j <= F_sp + (task_cost tsk - task_run_to_completion_threshold tsk)

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


              rewrite -addnBA // subnn addn0.

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

1 subgoal (ID 1014)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  PRT1 : task_rtc_bounded_by_cost tsk
  PRT2 : job_respects_task_rtc arr_seq tsk
  ============================
  F_sp - task_run_to_completion_threshold tsk + job_cost j <=
  F_sp + (task_cost tsk - task_run_to_completion_threshold tsk)

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


              rewrite addnBA; last by apply PRT1.

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

1 subgoal (ID 1020)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  PRT1 : task_rtc_bounded_by_cost tsk
  PRT2 : job_respects_task_rtc arr_seq tsk
  ============================
  F_sp - task_run_to_completion_threshold tsk + job_cost j <=
  F_sp + task_cost tsk - task_run_to_completion_threshold tsk

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


              rewrite subh1; last by done.

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

1 subgoal (ID 1027)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  PRT1 : task_rtc_bounded_by_cost tsk
  PRT2 : job_respects_task_rtc arr_seq tsk
  ============================
  F_sp + job_cost j - task_run_to_completion_threshold tsk <=
  F_sp + task_cost tsk - task_run_to_completion_threshold tsk

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


              rewrite leq_sub2r // leq_add2l.

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

1 subgoal (ID 1062)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : task_run_to_completion_threshold tsk <= F_sp
  L2 : optimism <= F
  PRT1 : task_rtc_bounded_by_cost tsk
  PRT2 : job_respects_task_rtc arr_seq tsk
  ============================
  job_cost j <= task_cost tsk

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


                by rewrite -H_job_of_tsk; apply H_valid_job_cost.

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

No more subgoals.

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


            }

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

No more subgoals.

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


          Qed.

... which contradicts the initial assumption about [j] is not completed by time [job_arrival j + R].
          Lemma j_is_completed_earlier_contradiction : False.

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

1 subgoal (ID 730)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  ============================
  False

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


          Proof.
            move: H_j_not_completed ⇒ /negP C; apply: C.

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

1 subgoal (ID 763)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  F : duration
  H_Asp_Fsp_eq_A_F : A_sp + F_sp = A + F
  H_F_le_Fsp : F <= F_sp
  H_A_F_fixpoint : A + F =
                   task_run_to_completion_threshold tsk +
                   interference_bound_function tsk A (A + F)
  H_j_not_completed : ~~ completed_by sched j (job_arrival j + R)
  job_last := job_cost j - job_run_to_completion_threshold j : nat
  optimism := task_run_to_completion_threshold tsk -
              job_run_to_completion_threshold j : nat
  ============================
  completed_by sched j (job_arrival j + R)

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


            apply completion_monotonic with (t1 + ((A + F) - optimism) + job_last);
              auto using j_is_completed_by_t1_A_F_optimist_last, t1_A_F_optimist_last_le_arrival_R.
(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

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


          Qed.

        End ProofByContradiction.

Putting everything together, we conclude that [j] is completed by [job_arrival j + R].
        Lemma job_completed_by_arrival_plus_R_2:
          completed_by sched j (job_arrival j + R).

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

1 subgoal (ID 706)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  ============================
  completed_by sched j (job_arrival j + R)

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


        Proof.
          have HelpAuto: m n, n n + m; first by intros; rewrite leq_addr.

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

1 subgoal (ID 711)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  ============================
  completed_by sched j (job_arrival j + R)

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


          move: H_busy_interval ⇒ [[/andP [GT LT] _] _].

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

1 subgoal (ID 778)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  completed_by sched j (job_arrival j + R)

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


          have L1 := solution_for_A_exists
                       tsk L (fun tsk A Rtask_run_to_completion_threshold tsk
                                          + interference_bound_function tsk A R) A_sp F_sp.

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

1 subgoal (ID 789)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : forall t : TaskRunToCompletionThreshold Task,
       A_sp + F_sp < L ->
       A_sp + F_sp =
       task_run_to_completion_threshold tsk +
       interference_bound_function tsk A_sp (A_sp + F_sp) ->
       forall A : duration,
       A_sp <= A <= A_sp + F_sp ->
       are_equivalent_at_values_less_than
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A R)
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A_sp R) L ->
       exists F : nat,
         A_sp + F_sp = A + F /\
         F <= F_sp /\
         A + F =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A (A + F)
  ============================
  completed_by sched j (job_arrival j + R)

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


          specialize (L1 H0).

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

1 subgoal (ID 791)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : A_sp + F_sp < L ->
       A_sp + F_sp =
       task_run_to_completion_threshold tsk +
       interference_bound_function tsk A_sp (A_sp + F_sp) ->
       forall A : duration,
       A_sp <= A <= A_sp + F_sp ->
       are_equivalent_at_values_less_than
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A R)
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A_sp R) L ->
       exists F : nat,
         A_sp + F_sp = A + F /\
         F <= F_sp /\
         A + F =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A (A + F)
  ============================
  completed_by sched j (job_arrival j + R)

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


          feed_n 2 L1; try done.

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

2 subgoals (ID 792)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : A_sp + F_sp < L ->
       A_sp + F_sp =
       task_run_to_completion_threshold tsk +
       interference_bound_function tsk A_sp (A_sp + F_sp) ->
       forall A : duration,
       A_sp <= A <= A_sp + F_sp ->
       are_equivalent_at_values_less_than
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A R)
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A_sp R) L ->
       exists F : nat,
         A_sp + F_sp = A + F /\
         F <= F_sp /\
         A + F =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A (A + F)
  ============================
  A_sp + F_sp < L

subgoal 2 (ID 803) is:
 completed_by sched j (job_arrival j + R)

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


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

1 subgoal (ID 792)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : A_sp + F_sp < L ->
       A_sp + F_sp =
       task_run_to_completion_threshold tsk +
       interference_bound_function tsk A_sp (A_sp + F_sp) ->
       forall A : duration,
       A_sp <= A <= A_sp + F_sp ->
       are_equivalent_at_values_less_than
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A R)
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A_sp R) L ->
       exists F : nat,
         A_sp + F_sp = A + F /\
         F <= F_sp /\
         A + F =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A (A + F)
  ============================
  A_sp + F_sp < L

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


move: (H_busy_interval_exists j H_j_arrives H_job_of_tsk H_job_cost_positive)
            ⇒ [t1' [t2' [_ [BOUND BUSY]]]].

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

1 subgoal (ID 890)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : A_sp + F_sp < L ->
       A_sp + F_sp =
       task_run_to_completion_threshold tsk +
       interference_bound_function tsk A_sp (A_sp + F_sp) ->
       forall A : duration,
       A_sp <= A <= A_sp + F_sp ->
       are_equivalent_at_values_less_than
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A R)
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A_sp R) L ->
       exists F : nat,
         A_sp + F_sp = A + F /\
         F <= F_sp /\
         A + F =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A (A + F)
  t1', t2' : nat
  BOUND : t2' <= t1' + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1' t2'
  ============================
  A_sp + F_sp < L

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


            have EQ:= busy_interval_is_unique _ _ _ _ _ _ _ _ H_busy_interval BUSY.

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

1 subgoal (ID 916)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : A_sp + F_sp < L ->
       A_sp + F_sp =
       task_run_to_completion_threshold tsk +
       interference_bound_function tsk A_sp (A_sp + F_sp) ->
       forall A : duration,
       A_sp <= A <= A_sp + F_sp ->
       are_equivalent_at_values_less_than
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A R)
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A_sp R) L ->
       exists F : nat,
         A_sp + F_sp = A + F /\
         F <= F_sp /\
         A + F =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A (A + F)
  t1', t2' : nat
  BOUND : t2' <= t1' + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1' t2'
  EQ : t1 = t1' /\ t2 = t2'
  ============================
  A_sp + F_sp < L

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


            destruct EQ as [EQ1 EQ2].

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

1 subgoal (ID 921)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : A_sp + F_sp < L ->
       A_sp + F_sp =
       task_run_to_completion_threshold tsk +
       interference_bound_function tsk A_sp (A_sp + F_sp) ->
       forall A : duration,
       A_sp <= A <= A_sp + F_sp ->
       are_equivalent_at_values_less_than
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A R)
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A_sp R) L ->
       exists F : nat,
         A_sp + F_sp = A + F /\
         F <= F_sp /\
         A + F =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A (A + F)
  t1', t2' : nat
  BOUND : t2' <= t1' + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1' t2'
  EQ1 : t1 = t1'
  EQ2 : t2 = t2'
  ============================
  A_sp + F_sp < L

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


            subst t1' t2'; clear BUSY.

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

1 subgoal (ID 938)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : A_sp + F_sp < L ->
       A_sp + F_sp =
       task_run_to_completion_threshold tsk +
       interference_bound_function tsk A_sp (A_sp + F_sp) ->
       forall A : duration,
       A_sp <= A <= A_sp + F_sp ->
       are_equivalent_at_values_less_than
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A R)
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A_sp R) L ->
       exists F : nat,
         A_sp + F_sp = A + F /\
         F <= F_sp /\
         A + F =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A (A + F)
  BOUND : t2 <= t1 + L
  ============================
  A_sp + F_sp < L

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


              by rewrite -(ltn_add2l t1); apply leq_trans with t2.

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

1 subgoal (ID 803)

subgoal 1 (ID 803) is:
 completed_by sched j (job_arrival j + R)

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


          }

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

1 subgoal (ID 803)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : forall A : duration,
       A_sp <= A <= A_sp + F_sp ->
       are_equivalent_at_values_less_than
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A R)
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A_sp R) L ->
       exists F : nat,
         A_sp + F_sp = A + F /\
         F <= F_sp /\
         A + F =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A (A + F)
  ============================
  completed_by sched j (job_arrival j + R)

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


          specialize (L1 A); feed_n 2 L1.

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

3 subgoals (ID 947)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : A_sp <= A <= A_sp + F_sp ->
       are_equivalent_at_values_less_than
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A R)
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A_sp R) L ->
       exists F : nat,
         A_sp + F_sp = A + F /\
         F <= F_sp /\
         A + F =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A (A + F)
  ============================
  A_sp <= A <= A_sp + F_sp

subgoal 2 (ID 953) is:
 are_equivalent_at_values_less_than
   (fun R0 : duration =>
    task_run_to_completion_threshold tsk +
    interference_bound_function tsk A R0)
   (fun R0 : duration =>
    task_run_to_completion_threshold tsk +
    interference_bound_function tsk A_sp R0) L
subgoal 3 (ID 958) is:
 completed_by sched j (job_arrival j + R)

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


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

1 subgoal (ID 947)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : A_sp <= A <= A_sp + F_sp ->
       are_equivalent_at_values_less_than
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A R)
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A_sp R) L ->
       exists F : nat,
         A_sp + F_sp = A + F /\
         F <= F_sp /\
         A + F =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A (A + F)
  ============================
  A_sp <= A <= A_sp + F_sp

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


by apply/andP; split.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 953)

subgoal 1 (ID 953) is:
 are_equivalent_at_values_less_than
   (fun R0 : duration =>
    task_run_to_completion_threshold tsk +
    interference_bound_function tsk A R0)
   (fun R0 : duration =>
    task_run_to_completion_threshold tsk +
    interference_bound_function tsk A_sp R0) L
subgoal 2 (ID 958) is:
 completed_by sched j (job_arrival j + R)

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


}

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

2 subgoals (ID 953)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : are_equivalent_at_values_less_than
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A R)
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A_sp R) L ->
       exists F : nat,
         A_sp + F_sp = A + F /\
         F <= F_sp /\
         A + F =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A (A + F)
  ============================
  are_equivalent_at_values_less_than
    (fun R0 : duration =>
     task_run_to_completion_threshold tsk +
     interference_bound_function tsk A R0)
    (fun R0 : duration =>
     task_run_to_completion_threshold tsk +
     interference_bound_function tsk A_sp R0) L

subgoal 2 (ID 958) is:
 completed_by sched j (job_arrival j + R)

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


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

1 subgoal (ID 953)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : are_equivalent_at_values_less_than
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A R)
         (fun R : duration =>
          task_run_to_completion_threshold tsk +
          interference_bound_function tsk A_sp R) L ->
       exists F : nat,
         A_sp + F_sp = A + F /\
         F <= F_sp /\
         A + F =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A (A + F)
  ============================
  are_equivalent_at_values_less_than
    (fun R0 : duration =>
     task_run_to_completion_threshold tsk +
     interference_bound_function tsk A R0)
    (fun R0 : duration =>
     task_run_to_completion_threshold tsk +
     interference_bound_function tsk A_sp R0) L

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


by intros x LTG; apply/eqP; rewrite eqn_add2l H_equivalent.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 958)

subgoal 1 (ID 958) is:
 completed_by sched j (job_arrival j + R)

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


}

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

1 subgoal (ID 958)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  L1 : exists F : nat,
         A_sp + F_sp = A + F /\
         F <= F_sp /\
         A + F =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A (A + F)
  ============================
  completed_by sched j (job_arrival j + R)

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


          move: L1 ⇒ [F [EQSUM [F2LEF1 FIX2]]].

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

1 subgoal (ID 1088)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  F : nat
  EQSUM : A_sp + F_sp = A + F
  F2LEF1 : F <= F_sp
  FIX2 : A + F =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A (A + F)
  ============================
  completed_by sched j (job_arrival j + R)

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


          apply/negP; intros CONTRc; move: CONTRc ⇒ /negP CONTRc.

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

1 subgoal (ID 1173)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_no_less_than_relative_arrival_of_j : A <= A_sp + F_sp
  HelpAuto : forall m n : nat, n <= n + m
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  F : nat
  EQSUM : A_sp + F_sp = A + F
  F2LEF1 : F <= F_sp
  FIX2 : A + F =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A (A + F)
  CONTRc : ~~ completed_by sched j (job_arrival j + R)
  ============================
  False

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


            by eapply j_is_completed_earlier_contradiction in CONTRc; eauto 2.

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

No more subgoals.

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


        Qed.

      End FixpointIsNoLessThanArrival.

CASE 2: the value of the fix-point is less than the relative arrival time of job j (which turns out to be impossible, i.e. the solution of the response-time recurrence is always equal to or greater than the relative arrival time).
      Section FixpointCannotBeSmallerThanArrival.

Assume that [A_sp + F_sp] is less than A.
Note that the relative arrival time of job j is less than L.
        Lemma relative_arrival_is_bounded: A < L.

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

1 subgoal (ID 700)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  ============================
  A < L

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


        Proof.
          rewrite /A.

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

1 subgoal (ID 701)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  ============================
  job_arrival j - t1 < L

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


          move: (H_busy_interval_exists j H_j_arrives H_job_of_tsk H_job_cost_positive) ⇒ [t1' [t2' [_ [BOUND BUSY]]]].

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

1 subgoal (ID 744)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  t1', t2' : nat
  BOUND : t2' <= t1' + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1' t2'
  ============================
  job_arrival j - t1 < L

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


          have EQ:= busy_interval_is_unique _ _ _ _ _ _ _ _ H_busy_interval BUSY.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 770)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  t1', t2' : nat
  BOUND : t2' <= t1' + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1' t2'
  EQ : t1 = t1' /\ t2 = t2'
  ============================
  job_arrival j - t1 < L

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


destruct EQ as [EQ1 EQ2].

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

1 subgoal (ID 775)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  t1', t2' : nat
  BOUND : t2' <= t1' + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1' t2'
  EQ1 : t1 = t1'
  EQ2 : t2 = t2'
  ============================
  job_arrival j - t1 < L

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


          subst t1' t2'; clear BUSY.

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

1 subgoal (ID 792)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  BOUND : t2 <= t1 + L
  ============================
  job_arrival j - t1 < L

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


          apply leq_trans with (t2 - t1); last by rewrite leq_subLR.

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

1 subgoal (ID 793)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  BOUND : t2 <= t1 + L
  ============================
  job_arrival j - t1 < t2 - t1

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


          move: (H_busy_interval)=> [[/andP [T1 T3] [_ _]] _].

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

1 subgoal (ID 870)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  BOUND : t2 <= t1 + L
  T1 : t1 <= job_arrival j
  T3 : job_arrival j < t2
  ============================
  job_arrival j - t1 < t2 - t1

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


            by apply ltn_sub2r; first apply leq_ltn_trans with (job_arrival j).

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

No more subgoals.

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


        Qed.

We can use [j_receives_at_least_run_to_completion_threshold] to prove that the service received by j by time [t1 + (A_sp + F_sp)] is no less than run-to-completion threshold.
        Lemma service_of_job_ge_run_to_completion_threshold:
          service sched j (t1 + (A_sp + F_sp)) job_run_to_completion_threshold j.

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

1 subgoal (ID 707)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  ============================
  job_run_to_completion_threshold j <= service sched j (t1 + (A_sp + F_sp))

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


        Proof.
          move: (H_busy_interval) ⇒ [[NEQ [QT1 NTQ]] QT2].

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

1 subgoal (ID 739)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  ============================
  job_run_to_completion_threshold j <= service sched j (t1 + (A_sp + F_sp))

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


          move: (NEQ) ⇒ /andP [GT LT].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 780)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  job_run_to_completion_threshold j <= service sched j (t1 + (A_sp + F_sp))

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


          move: (H_job_interference_is_bounded t1 t2 (A_sp + F_sp) j) ⇒ IB.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 782)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  IB : arrives_in arr_seq j ->
       job_task j = tsk ->
       definitions.busy_interval sched interference interfering_workload j t1
         t2 ->
       t1 + (A_sp + F_sp) < t2 ->
       ~~ completed_by sched j (t1 + (A_sp + F_sp)) ->
       let offset := job_arrival j - t1 in
       definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk offset (A_sp + F_sp)
  ============================
  job_run_to_completion_threshold j <= service sched j (t1 + (A_sp + F_sp))

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


          feed_n 5 IB; try done.

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

2 subgoals (ID 807)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  IB : ~~ completed_by sched j (t1 + (A_sp + F_sp)) ->
       let offset := job_arrival j - t1 in
       definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk offset (A_sp + F_sp)
  ============================
  ~~ completed_by sched j (t1 + (A_sp + F_sp))

subgoal 2 (ID 812) is:
 job_run_to_completion_threshold j <= service sched j (t1 + (A_sp + F_sp))

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


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

1 subgoal (ID 807)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  IB : ~~ completed_by sched j (t1 + (A_sp + F_sp)) ->
       let offset := job_arrival j - t1 in
       definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk offset (A_sp + F_sp)
  ============================
  ~~ completed_by sched j (t1 + (A_sp + F_sp))

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


apply/negP; intros COMPL.

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

1 subgoal (ID 878)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  IB : ~~ completed_by sched j (t1 + (A_sp + F_sp)) ->
       let offset := job_arrival j - t1 in
       definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk offset (A_sp + F_sp)
  COMPL : completed_by sched j (t1 + (A_sp + F_sp))
  ============================
  False

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


            apply completion_monotonic with (t' := t1 + A) in COMPL; try done; last first.

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

2 subgoals (ID 885)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  IB : ~~ completed_by sched j (t1 + (A_sp + F_sp)) ->
       let offset := job_arrival j - t1 in
       definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk offset (A_sp + F_sp)
  COMPL : completed_by sched j (t1 + (A_sp + F_sp))
  ============================
  t1 + (A_sp + F_sp) <= t1 + A

subgoal 2 (ID 883) is:
 False

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


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

1 subgoal (ID 885)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  IB : ~~ completed_by sched j (t1 + (A_sp + F_sp)) ->
       let offset := job_arrival j - t1 in
       definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk offset (A_sp + F_sp)
  COMPL : completed_by sched j (t1 + (A_sp + F_sp))
  ============================
  t1 + (A_sp + F_sp) <= t1 + A

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


by rewrite leq_add2l; apply ltnW.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 883)

subgoal 1 (ID 883) is:
 False
subgoal 2 (ID 812) is:
 job_run_to_completion_threshold j <= service sched j (t1 + (A_sp + F_sp))

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


}

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

1 subgoal (ID 883)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  IB : ~~ completed_by sched j (t1 + (A_sp + F_sp)) ->
       let offset := job_arrival j - t1 in
       definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk offset (A_sp + F_sp)
  COMPL : completed_by sched j (t1 + A)
  ============================
  False

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


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

1 subgoal (ID 883)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  IB : ~~ completed_by sched j (t1 + (A_sp + F_sp)) ->
       let offset := job_arrival j - t1 in
       definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk offset (A_sp + F_sp)
  COMPL : completed_by sched j (t1 + A)
  ============================
  False

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


rewrite /A subnKC in COMPL; last by done.

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

1 subgoal (ID 1020)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  IB : ~~ completed_by sched j (t1 + (A_sp + F_sp)) ->
       let offset := job_arrival j - t1 in
       definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk offset (A_sp + F_sp)
  COMPL : completed_by sched j (job_arrival j)
  ============================
  False

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


              move: COMPL; rewrite /completed_by leqNgt; move ⇒ /negP COMPL; apply: COMPL.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1134)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  IB : ~~ completed_by sched j (t1 + (A_sp + F_sp)) ->
       let offset := job_arrival j - t1 in
       definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk offset (A_sp + F_sp)
  ============================
  service sched j (job_arrival j) < job_cost j

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


              rewrite /service -(service_during_cat _ _ _ (job_arrival j)); last by apply/andP; split.

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

1 subgoal (ID 1167)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  IB : ~~ completed_by sched j (t1 + (A_sp + F_sp)) ->
       let offset := job_arrival j - t1 in
       definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk offset (A_sp + F_sp)
  ============================
  service_during sched j 0 (job_arrival j) +
  service_during sched j (job_arrival j) (job_arrival j) < 
  job_cost j

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


              rewrite (cumulative_service_before_job_arrival_zero) // add0n.

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

1 subgoal (ID 1239)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  IB : ~~ completed_by sched j (t1 + (A_sp + F_sp)) ->
       let offset := job_arrival j - t1 in
       definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk offset (A_sp + F_sp)
  ============================
  service_during sched j (job_arrival j) (job_arrival j) < job_cost j

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


                by rewrite /service_during big_geq //.

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

1 subgoal (ID 812)

subgoal 1 (ID 812) is:
 job_run_to_completion_threshold j <= service sched j (t1 + (A_sp + F_sp))

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


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

1 subgoal (ID 812)

subgoal 1 (ID 812) is:
 job_run_to_completion_threshold j <= service sched j (t1 + (A_sp + F_sp))

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



          }

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

1 subgoal (ID 812)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  IB : let offset := job_arrival j - t1 in
       definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk offset (A_sp + F_sp)
  ============================
  job_run_to_completion_threshold j <= service sched j (t1 + (A_sp + F_sp))

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


          rewrite -/A in IB.

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

1 subgoal (ID 1331)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  IB : let offset := A in
       definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk offset (A_sp + F_sp)
  ============================
  job_run_to_completion_threshold j <= service sched j (t1 + (A_sp + F_sp))

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


          have ALTT := relative_arrival_is_bounded.

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

1 subgoal (ID 1336)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  IB : let offset := A in
       definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk offset (A_sp + F_sp)
  ALTT : A < L
  ============================
  job_run_to_completion_threshold j <= service sched j (t1 + (A_sp + F_sp))

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


          simpl in IB; rewrite H_equivalent in IB; last by apply ltn_trans with A.

(* ----------------------------------[ 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
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ALTT : A < L
  IB : definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk A_sp (A_sp + F_sp)
  ============================
  job_run_to_completion_threshold j <= service sched j (t1 + (A_sp + F_sp))

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


          have ESERV :=
            @j_receives_at_least_run_to_completion_threshold
              _ _ H1 H2 H3 PState H5 _ _ arr_seq sched tsk
              interference interfering_workload _ j _ _ _ t1 t2 _ (job_run_to_completion_threshold j) _ (A_sp + F_sp).
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 1517)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ALTT : A < L
  IB : definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk A_sp (A_sp + F_sp)
  ESERV : ideal_progress_proc_model PState ->
          unit_service_proc_model PState ->
          definitions.work_conserving arr_seq sched tsk interference
            interfering_workload ->
          arrives_in arr_seq j ->
          job_task j = tsk ->
          job_cost_positive j ->
          definitions.busy_interval sched interference interfering_workload j
            t1 t2 ->
          forall (j0 : JobCost Job) (j1 : JobPreemptable Job),
          job_run_to_completion_threshold j <= job_cost j ->
          job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A_sp + F_sp)) <= A_sp + F_sp ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A_sp + F_sp))
  ============================
  job_run_to_completion_threshold j <= service sched j (t1 + (A_sp + F_sp))

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


          feed_n 7 ESERV; eauto 2.

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

1 subgoal (ID 1559)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ALTT : A < L
  IB : definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk A_sp (A_sp + F_sp)
  ESERV : forall (j0 : JobCost Job) (j1 : JobPreemptable Job),
          job_run_to_completion_threshold j <= job_cost j ->
          job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A_sp + F_sp)) <= A_sp + F_sp ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A_sp + F_sp))
  ============================
  job_run_to_completion_threshold j <= service sched j (t1 + (A_sp + F_sp))

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


          specialize (ESERV H3 H4).

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

1 subgoal (ID 1582)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ALTT : A < L
  IB : definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk A_sp (A_sp + F_sp)
  ESERV : job_run_to_completion_threshold j <= job_cost j ->
          job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A_sp + F_sp)) <= A_sp + F_sp ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A_sp + F_sp))
  ============================
  job_run_to_completion_threshold j <= service sched j (t1 + (A_sp + F_sp))

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


          feed_n 2 ESERV; eauto using job_run_to_completion_threshold_le_job_cost.

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

1 subgoal (ID 1589)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ALTT : A < L
  IB : definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
       interference_bound_function tsk A_sp (A_sp + F_sp)
  ESERV : job_run_to_completion_threshold j +
          definitions.cumul_interference interference j t1
            (t1 + (A_sp + F_sp)) <= A_sp + F_sp ->
          job_run_to_completion_threshold j <=
          service sched j (t1 + (A_sp + F_sp))
  ============================
  job_run_to_completion_threshold j +
  definitions.cumul_interference interference j t1 (t1 + (A_sp + F_sp)) <=
  A_sp + F_sp

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


            by rewrite {2}H_Asp_Fsp_fixpoint leq_add //; apply H_valid_run_to_completion_threshold.

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

No more subgoals.

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


        Qed.

However, this is a contradiction. Since job [j] has not yet arrived, its service is equal to [0]. However, run-to-completion threshold is always positive.
        Lemma relative_arrival_time_is_no_less_than_fixpoint:
          False.

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

1 subgoal (ID 708)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  ============================
  False

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


        Proof.
          move: (H_busy_interval) ⇒ [[NEQ [QT1 NTQ]] QT2].

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

1 subgoal (ID 740)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  ============================
  False

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


          move: (NEQ) ⇒ /andP [GT LT].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 781)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  False

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



          have ESERV := service_of_job_ge_run_to_completion_threshold.

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

1 subgoal (ID 786)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ESERV : job_run_to_completion_threshold j <=
          service sched j (t1 + (A_sp + F_sp))
  ============================
  False

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


          move: ESERV; rewrite leqNgt; move ⇒ /negP ESERV; apply: ESERV.

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

1 subgoal (ID 824)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  service sched j (t1 + (A_sp + F_sp)) < job_run_to_completion_threshold j

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


          rewrite /service cumulative_service_before_job_arrival_zero;
            eauto 5 using job_run_to_completion_threshold_positive.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 848)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  t1 + (A_sp + F_sp) <= job_arrival j

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


          rewrite -[X in _ X](@subnKC t1) //.

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

1 subgoal (ID 7617)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  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
  t1, t2 : instant
  H_busy_interval : busy_interval j t1 t2
  A := job_arrival j - t1 : nat
  A_sp, F_sp : duration
  H_A_gt_Asp : A_sp <= A
  H_equivalent : are_equivalent_at_values_less_than
                   (interference_bound_function tsk A)
                   (interference_bound_function tsk A_sp) L
  H_Asp_is_in_search_space : is_in_search_space A_sp
  H_Asp_Fsp_fixpoint : A_sp + F_sp =
                       task_run_to_completion_threshold tsk +
                       interference_bound_function tsk A_sp (A_sp + F_sp)
  H_R_gt_Fsp : F_sp + (task_cost tsk - task_run_to_completion_threshold tsk) <=
               R
  H_small_fixpoint_solution : t1 + (A_sp + F_sp) < t2
  H_fixpoint_is_less_that_relative_arrival_of_j : A_sp + F_sp < A
  NEQ : t1 <= job_arrival j < t2
  QT1 : quiet_time sched interference interfering_workload j t1
  NTQ : forall t : nat,
        t1 < t < t2 ->
        ~ quiet_time sched interference interfering_workload j t
  QT2 : quiet_time sched interference interfering_workload j t2
  GT : t1 <= job_arrival j
  LT : job_arrival j < t2
  ============================
  t1 + (A_sp + F_sp) <= t1 + (job_arrival j - t1)

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


            by rewrite -/A leq_add2l ltnW.

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

No more subgoals.

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


        Qed.

      End FixpointCannotBeSmallerThanArrival.

    End FixpointInsideBusyInterval.

  End ProofOfTheorem.

Using the lemmas above, we prove that [R] is a response-time bound.
  Theorem uniprocessor_response_time_bound:
    response_time_bounded_by tsk R.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 681)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  ============================
  response_time_bounded_by tsk R

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


  Proof.
    intros j ARR JOBtsk.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 685)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  ============================
  job_response_time_bound sched j R

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


unfold job_response_time_bound.

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

1 subgoal (ID 687)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  ============================
  completed_by sched j (job_arrival j + R)

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


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

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

2 subgoals (ID 704)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  ZERO : job_cost j = 0
  ============================
  completed_by sched j (job_arrival j + R)

subgoal 2 (ID 705) is:
 completed_by sched j (job_arrival j + R)

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


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

1 subgoal (ID 704)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  ZERO : job_cost j = 0
  ============================
  completed_by sched j (job_arrival j + R)

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


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

1 subgoal (ID 705)

subgoal 1 (ID 705) is:
 completed_by sched j (job_arrival j + R)

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


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

1 subgoal (ID 705)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  POS : 0 < job_cost j
  ============================
  completed_by sched j (job_arrival j + R)

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



    move: (H_busy_interval_exists j ARR JOBtsk POS) ⇒ [t1 [t2 [NEQ [T2 BUSY]]]].

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

1 subgoal (ID 756)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  POS : 0 < job_cost j
  t1, t2 : nat
  NEQ : t1 <= job_arrival j < t2
  T2 : t2 <= t1 + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1 t2
  ============================
  completed_by sched j (job_arrival j + R)

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


    move: (NEQ) (BUSY)=> /andP [GE LT] [_ QTt2].

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

1 subgoal (ID 809)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  POS : 0 < job_cost j
  t1, t2 : nat
  NEQ : t1 <= job_arrival j < t2
  T2 : t2 <= t1 + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1 t2
  GE : t1 <= job_arrival j
  LT : job_arrival j < t2
  QTt2 : quiet_time sched interference interfering_workload j t2
  ============================
  completed_by sched j (job_arrival j + R)

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


    have A2LTL := relative_arrival_is_bounded _ ARR JOBtsk POS _ _ BUSY.

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

1 subgoal (ID 820)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  POS : 0 < job_cost j
  t1, t2 : nat
  NEQ : t1 <= job_arrival j < t2
  T2 : t2 <= t1 + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1 t2
  GE : t1 <= job_arrival j
  LT : job_arrival j < t2
  QTt2 : quiet_time sched interference interfering_workload j t2
  A2LTL : job_arrival j - t1 < L
  ============================
  completed_by sched j (job_arrival j + R)

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


    set (A2 := job_arrival j - t1) in ×.

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

1 subgoal (ID 826)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  POS : 0 < job_cost j
  t1, t2 : nat
  NEQ : t1 <= job_arrival j < t2
  T2 : t2 <= t1 + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1 t2
  GE : t1 <= job_arrival j
  LT : job_arrival j < t2
  QTt2 : quiet_time sched interference interfering_workload j t2
  A2 := job_arrival j - t1 : nat
  A2LTL : A2 < L
  ============================
  completed_by sched j (job_arrival j + R)

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


    move: (representative_exists tsk _ interference_bound_function _ A2LTL) ⇒ [A1 [ALEA2 [EQΦ INSP]]].

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

1 subgoal (ID 863)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  POS : 0 < job_cost j
  t1, t2 : nat
  NEQ : t1 <= job_arrival j < t2
  T2 : t2 <= t1 + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1 t2
  GE : t1 <= job_arrival j
  LT : job_arrival j < t2
  QTt2 : quiet_time sched interference interfering_workload j t2
  A2 := job_arrival j - t1 : nat
  A2LTL : A2 < L
  A1 : nat
  ALEA2 : A1 <= A2
  EQΦ : are_equivalent_at_values_less_than
          (interference_bound_function tsk A2)
          (interference_bound_function tsk A1) L
  INSP : search_space.is_in_search_space tsk L interference_bound_function A1
  ============================
  completed_by sched j (job_arrival j + R)

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


    move: (H_R_is_maximum _ INSP) ⇒ [F1 [FIX1 LE1]].

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

1 subgoal (ID 887)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  POS : 0 < job_cost j
  t1, t2 : nat
  NEQ : t1 <= job_arrival j < t2
  T2 : t2 <= t1 + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1 t2
  GE : t1 <= job_arrival j
  LT : job_arrival j < t2
  QTt2 : quiet_time sched interference interfering_workload j t2
  A2 := job_arrival j - t1 : nat
  A2LTL : A2 < L
  A1 : nat
  ALEA2 : A1 <= A2
  EQΦ : are_equivalent_at_values_less_than
          (interference_bound_function tsk A2)
          (interference_bound_function tsk A1) L
  INSP : search_space.is_in_search_space tsk L interference_bound_function A1
  F1 : nat
  FIX1 : A1 + F1 =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A1 (A1 + F1)
  LE1 : F1 + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
  ============================
  completed_by sched j (job_arrival j + R)

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


    destruct (t1 + (A1 + F1) t2) eqn:BIG.

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

2 subgoals (ID 897)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  POS : 0 < job_cost j
  t1, t2 : nat
  NEQ : t1 <= job_arrival j < t2
  T2 : t2 <= t1 + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1 t2
  GE : t1 <= job_arrival j
  LT : job_arrival j < t2
  QTt2 : quiet_time sched interference interfering_workload j t2
  A2 := job_arrival j - t1 : nat
  A2LTL : A2 < L
  A1 : nat
  ALEA2 : A1 <= A2
  EQΦ : are_equivalent_at_values_less_than
          (interference_bound_function tsk A2)
          (interference_bound_function tsk A1) L
  INSP : search_space.is_in_search_space tsk L interference_bound_function A1
  F1 : nat
  FIX1 : A1 + F1 =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A1 (A1 + F1)
  LE1 : F1 + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
  BIG : (t2 <= t1 + (A1 + F1)) = true
  ============================
  completed_by sched j (job_arrival j + R)

subgoal 2 (ID 898) is:
 completed_by sched j (job_arrival j + R)

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


    - eapply job_completed_by_arrival_plus_R_1; eauto 2.

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

1 subgoal (ID 898)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  POS : 0 < job_cost j
  t1, t2 : nat
  NEQ : t1 <= job_arrival j < t2
  T2 : t2 <= t1 + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1 t2
  GE : t1 <= job_arrival j
  LT : job_arrival j < t2
  QTt2 : quiet_time sched interference interfering_workload j t2
  A2 := job_arrival j - t1 : nat
  A2LTL : A2 < L
  A1 : nat
  ALEA2 : A1 <= A2
  EQΦ : are_equivalent_at_values_less_than
          (interference_bound_function tsk A2)
          (interference_bound_function tsk A1) L
  INSP : search_space.is_in_search_space tsk L interference_bound_function A1
  F1 : nat
  FIX1 : A1 + F1 =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A1 (A1 + F1)
  LE1 : F1 + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
  BIG : (t2 <= t1 + (A1 + F1)) = false
  ============================
  completed_by sched j (job_arrival j + R)

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


    - apply negbT in BIG; rewrite -ltnNge in BIG.

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

1 subgoal (ID 1000)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  POS : 0 < job_cost j
  t1, t2 : nat
  NEQ : t1 <= job_arrival j < t2
  T2 : t2 <= t1 + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1 t2
  GE : t1 <= job_arrival j
  LT : job_arrival j < t2
  QTt2 : quiet_time sched interference interfering_workload j t2
  A2 := job_arrival j - t1 : nat
  A2LTL : A2 < L
  A1 : nat
  ALEA2 : A1 <= A2
  EQΦ : are_equivalent_at_values_less_than
          (interference_bound_function tsk A2)
          (interference_bound_function tsk A1) L
  INSP : search_space.is_in_search_space tsk L interference_bound_function A1
  F1 : nat
  FIX1 : A1 + F1 =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A1 (A1 + F1)
  LE1 : F1 + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
  BIG : t1 + (A1 + F1) < t2
  ============================
  completed_by sched j (job_arrival j + R)

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


      destruct (A2 A1 + F1) eqn:BOUND.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 1010)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  POS : 0 < job_cost j
  t1, t2 : nat
  NEQ : t1 <= job_arrival j < t2
  T2 : t2 <= t1 + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1 t2
  GE : t1 <= job_arrival j
  LT : job_arrival j < t2
  QTt2 : quiet_time sched interference interfering_workload j t2
  A2 := job_arrival j - t1 : nat
  A2LTL : A2 < L
  A1 : nat
  ALEA2 : A1 <= A2
  EQΦ : are_equivalent_at_values_less_than
          (interference_bound_function tsk A2)
          (interference_bound_function tsk A1) L
  INSP : search_space.is_in_search_space tsk L interference_bound_function A1
  F1 : nat
  FIX1 : A1 + F1 =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A1 (A1 + F1)
  LE1 : F1 + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
  BIG : t1 + (A1 + F1) < t2
  BOUND : (A2 <= A1 + F1) = true
  ============================
  completed_by sched j (job_arrival j + R)

subgoal 2 (ID 1011) is:
 completed_by sched j (job_arrival j + R)

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


      + eapply job_completed_by_arrival_plus_R_2; eauto 2.

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

1 subgoal (ID 1011)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  POS : 0 < job_cost j
  t1, t2 : nat
  NEQ : t1 <= job_arrival j < t2
  T2 : t2 <= t1 + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1 t2
  GE : t1 <= job_arrival j
  LT : job_arrival j < t2
  QTt2 : quiet_time sched interference interfering_workload j t2
  A2 := job_arrival j - t1 : nat
  A2LTL : A2 < L
  A1 : nat
  ALEA2 : A1 <= A2
  EQΦ : are_equivalent_at_values_less_than
          (interference_bound_function tsk A2)
          (interference_bound_function tsk A1) L
  INSP : search_space.is_in_search_space tsk L interference_bound_function A1
  F1 : nat
  FIX1 : A1 + F1 =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A1 (A1 + F1)
  LE1 : F1 + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
  BIG : t1 + (A1 + F1) < t2
  BOUND : (A2 <= A1 + F1) = false
  ============================
  completed_by sched j (job_arrival j + R)

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


      + apply negbT in BOUND; rewrite -ltnNge in BOUND.

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

1 subgoal (ID 1132)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskRunToCompletionThreshold Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : JobPreemptable Job
  PState : Type
  H5 : ProcessorState Job PState
  H_ideal_progress_proc_model : ideal_progress_proc_model PState
  H_unit_service_proc_model : unit_service_proc_model PState
  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 PState
  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
  H_valid_job_cost : arrivals_have_valid_job_costs arr_seq
  ts : seq Task
  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
  work_conserving := definitions.work_conserving arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> Prop
  busy_intervals_are_bounded_by := definitions.busy_intervals_are_bounded_by
                                     arr_seq sched tsk
   : (Job -> instant -> bool) -> (Job -> instant -> duration) -> nat -> Prop
  job_interference_is_bounded_by := definitions.job_interference_is_bounded_by
                                      arr_seq sched tsk
   : (Job -> instant -> bool) ->
     (Job -> instant -> duration) ->
     (Task -> duration -> duration -> duration) -> Prop
  interference : Job -> instant -> bool
  interfering_workload : Job -> instant -> duration
  H_work_conserving : work_conserving interference interfering_workload
  cumul_interference := definitions.cumul_interference interference
   : Job -> nat -> nat -> nat
  cumul_interfering_workload := definitions.cumul_interfering_workload
                                  interfering_workload
   : Job -> nat -> nat -> nat
  busy_interval := definitions.busy_interval sched interference
                     interfering_workload : Job -> instant -> instant -> Prop
  response_time_bounded_by := task_response_time_bound arr_seq sched
   : Task -> duration -> Prop
  L : duration
  H_busy_interval_exists : busy_intervals_are_bounded_by interference
                             interfering_workload L
  interference_bound_function : Task -> duration -> duration -> duration
  H_job_interference_is_bounded : job_interference_is_bounded_by interference
                                    interfering_workload
                                    interference_bound_function
  is_in_search_space := [eta search_space.is_in_search_space tsk L
                               interference_bound_function] : 
  nat -> Prop
  R : nat
  H_R_is_maximum : forall A : nat,
                   is_in_search_space A ->
                   exists F : nat,
                     A + F =
                     task_run_to_completion_threshold tsk +
                     interference_bound_function tsk A (A + F) /\
                     F +
                     (task_cost tsk - task_run_to_completion_threshold tsk) <=
                     R
  j : Job
  ARR : arrives_in arr_seq j
  JOBtsk : job_task j = tsk
  POS : 0 < job_cost j
  t1, t2 : nat
  NEQ : t1 <= job_arrival j < t2
  T2 : t2 <= t1 + L
  BUSY : definitions.busy_interval sched interference interfering_workload j
           t1 t2
  GE : t1 <= job_arrival j
  LT : job_arrival j < t2
  QTt2 : quiet_time sched interference interfering_workload j t2
  A2 := job_arrival j - t1 : nat
  A2LTL : A2 < L
  A1 : nat
  ALEA2 : A1 <= A2
  EQΦ : are_equivalent_at_values_less_than
          (interference_bound_function tsk A2)
          (interference_bound_function tsk A1) L
  INSP : search_space.is_in_search_space tsk L interference_bound_function A1
  F1 : nat
  FIX1 : A1 + F1 =
         task_run_to_completion_threshold tsk +
         interference_bound_function tsk A1 (A1 + F1)
  LE1 : F1 + (task_cost tsk - task_run_to_completion_threshold tsk) <= R
  BIG : t1 + (A1 + F1) < t2
  BOUND : A1 + F1 < A2
  ============================
  completed_by sched j (job_arrival j + R)

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


        exfalso; apply relative_arrival_time_is_no_less_than_fixpoint
                   with (j := j) (t1 := t1) (t2 := t2) (A_sp := A1) (F_sp := F1); auto.

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

No more subgoals.

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


  Qed.

End Abstract_RTA.