Library prosa.analysis.facts.busy_interval.priority_inversion


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

Welcome to Coq 8.11.2 (June 2020)

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


Require Export prosa.model.task.preemption.parameters.
Require Export prosa.model.schedule.priority_driven.
Require Export prosa.model.schedule.work_conserving.
Require Export prosa.analysis.definitions.job_properties.
Require Export prosa.analysis.definitions.busy_interval.
Require Export prosa.analysis.facts.model.ideal_schedule.
Require Export prosa.analysis.facts.busy_interval.busy_interval.

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

Throughout the file we assume for the classic Liu & Layland model of readiness without jitter and no self-suspensions, where pending jobs are always ready.
Require Import prosa.model.readiness.basic.

In preparation of the derivation of the priority inversion bound, we establish two basic facts on preemption times.
Section PreemptionTimes.

Consider any type of tasks ...
  Context {Task : TaskType}.
  Context `{TaskCost Task}.
  Context `{TaskMaxNonpreemptiveSegment Task}.

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

In addition, we assume the existence of a function mapping a task to its maximal non-preemptive segment ...
... and the existence of a function mapping a job and its progress to a boolean value saying whether this job is preemptable at its current point of execution.
  Context `{JobPreemptable Job}.

Consider any arrival sequence with consistent arrivals.
Next, consider any ideal uniprocessor schedule of this arrival sequence ...
Consider a valid model with bounded nonpreemptive segments.
Then, we show that time 0 is a preemption time.
  Lemma zero_is_pt: preemption_time sched 0.

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

1 subgoal (ID 2431)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskMaxNonpreemptiveSegment Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_model_with_bounded_nonpreemptive_segments : valid_model_with_bounded_nonpreemptive_segments
                                                 arr_seq sched
  ============================
  preemption_time sched 0

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


  Proof.
    unfold preemption_time.

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

1 subgoal (ID 2433)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskMaxNonpreemptiveSegment Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_model_with_bounded_nonpreemptive_segments : valid_model_with_bounded_nonpreemptive_segments
                                                 arr_seq sched
  ============================
  match sched 0 with
  | Some j => job_preemptable j (service sched j 0)
  | None => true
  end

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


    case SCHED: (sched 0) ⇒ [j | ]; last by done.

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

1 subgoal (ID 2460)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskMaxNonpreemptiveSegment Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_model_with_bounded_nonpreemptive_segments : valid_model_with_bounded_nonpreemptive_segments
                                                 arr_seq sched
  j : Job
  SCHED : sched 0 = Some j
  ============================
  job_preemptable j (service sched j 0)

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


    move: (SCHED) ⇒ /eqP; rewrite -scheduled_at_def; moveARR.

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

1 subgoal (ID 2544)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskMaxNonpreemptiveSegment Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_model_with_bounded_nonpreemptive_segments : valid_model_with_bounded_nonpreemptive_segments
                                                 arr_seq sched
  j : Job
  SCHED : sched 0 = Some j
  ARR : scheduled_at sched j 0
  ============================
  job_preemptable j (service sched j 0)

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


    apply H_jobs_come_from_arrival_sequence in ARR.

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

1 subgoal (ID 2545)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskMaxNonpreemptiveSegment Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_model_with_bounded_nonpreemptive_segments : valid_model_with_bounded_nonpreemptive_segments
                                                 arr_seq sched
  j : Job
  SCHED : sched 0 = Some j
  ARR : arrives_in arr_seq j
  ============================
  job_preemptable j (service sched j 0)

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


    rewrite /service /service_during big_geq; last by done.

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

1 subgoal (ID 2570)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskMaxNonpreemptiveSegment Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_model_with_bounded_nonpreemptive_segments : valid_model_with_bounded_nonpreemptive_segments
                                                 arr_seq sched
  j : Job
  SCHED : sched 0 = Some j
  ARR : arrives_in arr_seq j
  ============================
  job_preemptable j 0

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


    destruct H_model_with_bounded_nonpreemptive_segments as [T1 T2].

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

1 subgoal (ID 2577)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskMaxNonpreemptiveSegment Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_model_with_bounded_nonpreemptive_segments : valid_model_with_bounded_nonpreemptive_segments
                                                 arr_seq sched
  j : Job
  SCHED : sched 0 = Some j
  ARR : arrives_in arr_seq j
  T1 : valid_preemption_model arr_seq sched
  T2 : model_with_bounded_nonpreemptive_segments arr_seq
  ============================
  job_preemptable j 0

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


    by move: (T1 j ARR) ⇒ [PP _].

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

No more subgoals.

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


  Qed.

Also, we show that the first instant of execution is a preemption time.
  Lemma first_moment_is_pt:
     j prt,
      arrives_in arr_seq j
      ~~ scheduled_at sched j prt
      scheduled_at sched j prt.+1
      preemption_time sched prt.+1.

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

1 subgoal (ID 2443)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskMaxNonpreemptiveSegment Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_model_with_bounded_nonpreemptive_segments : valid_model_with_bounded_nonpreemptive_segments
                                                 arr_seq sched
  ============================
  forall (j : Job) (prt : instant),
  arrives_in arr_seq j ->
  ~~ scheduled_at sched j prt ->
  scheduled_at sched j (succn prt) -> preemption_time sched (succn prt)

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


  Proof.
    intros s pt ARR NSCHED SCHED.

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

1 subgoal (ID 2448)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskMaxNonpreemptiveSegment Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_model_with_bounded_nonpreemptive_segments : valid_model_with_bounded_nonpreemptive_segments
                                                 arr_seq sched
  s : Job
  pt : instant
  ARR : arrives_in arr_seq s
  NSCHED : ~~ scheduled_at sched s pt
  SCHED : scheduled_at sched s (succn pt)
  ============================
  preemption_time sched (succn pt)

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


    unfold preemption_time.

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

1 subgoal (ID 2450)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskMaxNonpreemptiveSegment Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_model_with_bounded_nonpreemptive_segments : valid_model_with_bounded_nonpreemptive_segments
                                                 arr_seq sched
  s : Job
  pt : instant
  ARR : arrives_in arr_seq s
  NSCHED : ~~ scheduled_at sched s pt
  SCHED : scheduled_at sched s (succn pt)
  ============================
  match sched (succn pt) with
  | Some j => job_preemptable j (service sched j (succn pt))
  | None => true
  end

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


    move: (SCHED); rewrite scheduled_at_def; move ⇒ /eqP SCHED2; rewrite SCHED2; clear SCHED2.

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

1 subgoal (ID 2495)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskMaxNonpreemptiveSegment Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_model_with_bounded_nonpreemptive_segments : valid_model_with_bounded_nonpreemptive_segments
                                                 arr_seq sched
  s : Job
  pt : instant
  ARR : arrives_in arr_seq s
  NSCHED : ~~ scheduled_at sched s pt
  SCHED : scheduled_at sched s (succn pt)
  ============================
  job_preemptable s (service sched s (succn pt))

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


    destruct H_model_with_bounded_nonpreemptive_segments as [T1 T2].

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

1 subgoal (ID 2501)
  
  Task : TaskType
  H : TaskCost Task
  H0 : TaskMaxNonpreemptiveSegment Task
  Job : JobType
  H1 : JobTask Job Task
  H2 : JobArrival Job
  H3 : JobCost Job
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_model_with_bounded_nonpreemptive_segments : valid_model_with_bounded_nonpreemptive_segments
                                                 arr_seq sched
  s : Job
  pt : instant
  ARR : arrives_in arr_seq s
  NSCHED : ~~ scheduled_at sched s pt
  SCHED : scheduled_at sched s (succn pt)
  T1 : valid_preemption_model arr_seq sched
  T2 : model_with_bounded_nonpreemptive_segments arr_seq
  ============================
  job_preemptable s (service sched s (succn pt))

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


    by move: (T1 s ARR) ⇒ [_ [_ [_ P]]]; apply P.

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

No more subgoals.

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


  Qed.

End PreemptionTimes.

Priority inversion is bounded

In this module we prove that any priority inversion that occurs in the model with bounded nonpreemptive segments defined in module prosa.model.schedule.uni.limited.platform.definitions is bounded.
Consider any type of tasks ...
  Context {Task : TaskType}.
  Context `{TaskCost Task}.

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

Consider any arrival sequence with consistent arrivals ...
... and any ideal uniprocessor schedule of this arrival sequence ...
... where jobs do not execute before their arrival or after completion.
Consider a JLFP policy that indicates a higher-or-equal priority relation, and assume that the relation is reflexive and transitive.
In addition, we assume the existence of a function mapping a task to its maximal non-preemptive segment ...
... and the existence of a function mapping a job and its progress to a boolean value saying whether this job is preemptable at its current point of execution.
  Context `{JobPreemptable Job}.

And assume that they define a valid preemption model with bounded nonpreemptive segments.
Next, we assume that the schedule is a work-conserving schedule...
... and the schedule respects the policy defined by the predicate [job_preemptable] (i.e., bounded nonpreemptive segments).
Let's define some local names for clarity.
Finally, we introduce the notion of the maximal length of (potential) priority inversion at a time instant [t], which is defined as the maximum length of nonpreemptive segments among all jobs that arrived so far. Note that the value [job_max_nonpreemptive_segment j_lp] is at least [ε] for any job [j_lp], so the maximal length of priority inversion cannot be negative.
Next we prove that a priority inversion of a job is bounded by function [max_length_of_priority_inversion].
Note that any bound on function [max_length_of_priority_inversion] will also be a bound on the maximal priority inversion. This bound may be different for different scheduler and/or task models. Thus, we don't define such a bound in this module.
Consider any job [j] of [tsk] with positive job cost.
  Variable j : Job.
  Hypothesis H_j_arrives : arrives_in arr_seq j.
  Hypothesis H_job_cost_positive : job_cost_positive j.

Consider any busy interval prefix [t1, t2) of job j.
  Variable t1 t2 : instant.
  Hypothesis H_busy_interval_prefix:
    busy_interval_prefix arr_seq sched j t1 t2.

Processor Executes HEP jobs after Preemption Point

In this section, we prove that at any time instant after any preemption point (inside the busy interval), the processor is always busy scheduling a job with higher or equal priority.
First, we show that the processor at any preemptive point is always busy scheduling a job with higher or equal priority.
Consider an arbitrary preemption time t ∈ [t1,t2).
      Variable t : instant.
      Hypothesis H_t_in_busy_interval : t1 t < t2.
      Hypothesis H_t_preemption_time : preemption_time sched t.

First note since [t] lies inside the busy interval, the processor cannot be idle at time [t].
      Lemma instant_t_is_not_idle:
        ¬ is_idle sched t.

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

1 subgoal (ID 2487)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  ============================
  ~ is_idle sched t

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


      Proof.
        intros Idle.

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

1 subgoal (ID 2489)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  Idle : is_idle sched t
  ============================
  False

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


          by exfalso; eapply not_quiet_implies_not_idle with (t0 := t); eauto 2.

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

No more subgoals.

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


      Qed.

Next we consider two cases: (1) when [t] is less than [t2 - 1] and (2) [t] is equal to [t2 - 1].
      Lemma t_lt_t2_or_t_eq_t2:
        t < t2.-1 t = t2.-1.

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

1 subgoal (ID 2490)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  ============================
  t < predn t2 \/ t = predn t2

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


      Proof.
        have TEMP: t t2.-1.

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

2 subgoals (ID 2491)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  ============================
  t <= predn t2

subgoal 2 (ID 2493) is:
 t < predn t2 \/ t = predn t2

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


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

1 subgoal (ID 2491)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  ============================
  t <= predn t2

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


move: (H_t_in_busy_interval) ⇒ /andP [GEt LEt].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2534)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  GEt : t1 <= t
  LEt : t < t2
  ============================
  t <= predn t2

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


            by rewrite -subn1 subh3 // addn1.

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

1 subgoal (ID 2493)

subgoal 1 (ID 2493) is:
 t < predn t2 \/ t = predn t2

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


        }

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

1 subgoal (ID 2493)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  TEMP : t <= predn t2
  ============================
  t < predn t2 \/ t = predn t2

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


        rewrite leq_eqVlt in TEMP; move: TEMP ⇒ /orP [/eqP EQUALt2m1 | LTt2m1].

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

2 subgoals (ID 2689)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  ============================
  t < predn t2 \/ t = predn t2

subgoal 2 (ID 2690) is:
 t < predn t2 \/ t = predn t2

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


        - right; auto.

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

1 subgoal (ID 2690)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  ============================
  t < predn t2 \/ t = predn t2

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


        - left; auto.

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

No more subgoals.

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


      Qed.

In case when [t < t2 - 1], we use the fact that time instant [t+1] is not a quiet time. The later implies that there is a pending higher-or-equal priority job at time [t]. Thus, the assumption that the schedule respects priority policy at preemption points implies that the scheduled job must have a higher-or-equal priority.
      Lemma scheduled_at_preemption_time_implies_higher_or_equal_priority_lt:
        t < t2.-1
         jhp,
          scheduled_at sched jhp t
          hep_job jhp j.

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

1 subgoal (ID 2499)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  ============================
  t < predn t2 -> forall jhp : Job, scheduled_at sched jhp t -> hep_job jhp j

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


      Proof.
        intros LTt2m1 jhp Sched_jhp.

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

1 subgoal (ID 2502)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  ============================
  hep_job jhp j

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


        move: (H_t_in_busy_interval) (H_busy_interval_prefix) ⇒ /andP [GEt LEt] [SL [QUIET [NOTQUIET INBI]]].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2575)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  ============================
  hep_job jhp j

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


        apply contraT; move ⇒ /negP NOTHP; exfalso.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2603)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  ============================
  False

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


        apply NOTQUIET with (t := t.+1).

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

2 subgoals (ID 2604)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  ============================
  t1 < succn t < t2

subgoal 2 (ID 2605) is:
 quiet_time arr_seq sched j (succn t)

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


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

1 subgoal (ID 2604)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  ============================
  t1 < succn t < t2

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


apply/andP; split.

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

2 subgoals (ID 2631)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  ============================
  t1 < succn t

subgoal 2 (ID 2632) is:
 succn t < t2

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


          - by apply leq_ltn_trans with t1.

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

1 subgoal (ID 2632)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  ============================
  succn t < t2

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


          - rewrite -subn1 ltn_subRL addnC in LTt2m1.

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

1 subgoal (ID 2696)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  LTt2m1 : t + 1 < t2
  ============================
  succn t < t2

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


              by rewrite -[t.+1]addn1.

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

1 subgoal (ID 2605)

subgoal 1 (ID 2605) is:
 quiet_time arr_seq sched j (succn t)

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


        }

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

1 subgoal (ID 2605)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  ============================
  quiet_time arr_seq sched j (succn t)

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


        intros j_hp' IN HP ARR.

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

1 subgoal (ID 2704)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' (succn t)
  ============================
  completed_by sched j_hp' (succn t)

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


        apply contraT; move ⇒ /negP NOTCOMP'; exfalso.

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

1 subgoal (ID 2732)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' (succn t)
  NOTCOMP' : ~ completed_by sched j_hp' (succn t)
  ============================
  False

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


        have BACK: backlogged sched j_hp' t.

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

2 subgoals (ID 2739)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' (succn t)
  NOTCOMP' : ~ completed_by sched j_hp' (succn t)
  ============================
  backlogged sched j_hp' t

subgoal 2 (ID 2741) is:
 False

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


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

1 subgoal (ID 2739)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' (succn t)
  NOTCOMP' : ~ completed_by sched j_hp' (succn t)
  ============================
  backlogged sched j_hp' t

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


apply/andP; split; last first.

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

2 subgoals (ID 2768)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' (succn t)
  NOTCOMP' : ~ completed_by sched j_hp' (succn t)
  ============================
  ~~ scheduled_at sched j_hp' t

subgoal 2 (ID 2767) is:
 job_ready sched j_hp' t

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


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

1 subgoal (ID 2768)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' (succn t)
  NOTCOMP' : ~ completed_by sched j_hp' (succn t)
  ============================
  ~~ scheduled_at sched j_hp' t

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


apply/negP; intro SCHED'.

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

1 subgoal (ID 2790)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' (succn t)
  NOTCOMP' : ~ completed_by sched j_hp' (succn t)
  SCHED' : scheduled_at sched j_hp' t
  ============================
  False

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


            move: (ideal_proc_model_is_a_uniprocessor_model jhp j_hp' sched t Sched_jhp SCHED') ⇒ EQ; subst.

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

1 subgoal (ID 2801)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  j_hp' : Job
  Sched_jhp : scheduled_at sched j_hp' t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job j_hp' j
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' (succn t)
  NOTCOMP' : ~ completed_by sched j_hp' (succn t)
  SCHED' : scheduled_at sched j_hp' t
  ============================
  False

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


              by apply: NOTHP.

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

2 subgoals (ID 2767)

subgoal 1 (ID 2767) is:
 job_ready sched j_hp' t
subgoal 2 (ID 2741) is:
 False

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


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

1 subgoal (ID 2767)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' (succn t)
  NOTCOMP' : ~ completed_by sched j_hp' (succn t)
  ============================
  job_ready sched j_hp' t

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



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

2 subgoals (ID 2833)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' (succn t)
  NOTCOMP' : ~ completed_by sched j_hp' (succn t)
  ============================
  has_arrived j_hp' t

subgoal 2 (ID 2834) is:
 ~~ completed_by sched j_hp' t

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


unfold arrived_before, has_arrived in ×.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2836)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : job_arrival j_hp' < succn t
  NOTCOMP' : ~ completed_by sched j_hp' (succn t)
  ============================
  job_arrival j_hp' <= t

subgoal 2 (ID 2834) is:
 ~~ completed_by sched j_hp' t

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


by done.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2834)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' (succn t)
  NOTCOMP' : ~ completed_by sched j_hp' (succn t)
  ============================
  ~~ completed_by sched j_hp' t

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


          apply/negP; intro COMP; apply NOTCOMP'.

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

1 subgoal (ID 2859)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' (succn t)
  NOTCOMP' : ~ completed_by sched j_hp' (succn t)
  COMP : completed_by sched j_hp' t
  ============================
  completed_by sched j_hp' (succn t)

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


            by apply completion_monotonic with (t0 := t).

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

1 subgoal (ID 2741)

subgoal 1 (ID 2741) is:
 False

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


        }

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

1 subgoal (ID 2741)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  LTt2m1 : t < predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' (succn t)
  NOTCOMP' : ~ completed_by sched j_hp' (succn t)
  BACK : backlogged sched j_hp' t
  ============================
  False

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


        apply NOTHP, (H_priority_is_transitive t j_hp'); eauto 2.

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

No more subgoals.

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


      Qed.

In case when [t = t2 - 1], we cannot use the same proof since [t+1 = t2], but [t2] is a quiet time. So we do a similar case analysis on the fact that [t1 = t ∨ t1 < t].
      Lemma scheduled_at_preemption_time_implies_higher_or_equal_priority_eq:
        t = t2.-1
         jhp,
          scheduled_at sched jhp t
          hep_job jhp j.

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

1 subgoal (ID 2510)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  ============================
  t = predn t2 -> forall jhp : Job, scheduled_at sched jhp t -> hep_job jhp j

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


      Proof.
        intros EQUALt2m1 jhp Sched_jhp.

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

1 subgoal (ID 2513)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  ============================
  hep_job jhp j

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


        move: (H_t_in_busy_interval) (H_busy_interval_prefix) ⇒ /andP [GEt LEt] [SL [QUIET [NOTQUIET INBI]]].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2586)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  ============================
  hep_job jhp j

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


        apply contraT; move ⇒ /negP NOTHP; exfalso.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2614)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  ============================
  False

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


        rewrite leq_eqVlt in GEt; first move: GEt ⇒ /orP [/eqP EQUALt1 | LARGERt1].

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

2 subgoals (ID 2743)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  EQUALt1 : t1 = t
  ============================
  False

subgoal 2 (ID 2744) is:
 False

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


        - subst t; clear LEt.

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

2 subgoals (ID 2759)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp (predn t2)
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  EQUALt1 : t1 = predn t2
  ============================
  False

subgoal 2 (ID 2744) is:
 False

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


          rewrite -EQUALt1 in Sched_jhp; move: EQUALt1 ⇒ /eqP EQUALt1.

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

2 subgoals (ID 2867)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  EQUALt1 : t1 == predn t2
  ============================
  False

subgoal 2 (ID 2744) is:
 False

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


          destruct (job_scheduled_at j t1) eqn:SCHEDj.

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

3 subgoals (ID 2877)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  EQUALt1 : t1 == predn t2
  SCHEDj : job_scheduled_at j t1 = true
  ============================
  False

subgoal 2 (ID 2878) is:
 False
subgoal 3 (ID 2744) is:
 False

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


          + apply NOTHP; erewrite (ideal_proc_model_is_a_uniprocessor_model j jhp); eauto.

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

3 subgoals (ID 2883)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  EQUALt1 : t1 == predn t2
  SCHEDj : job_scheduled_at j t1 = true
  ============================
  hep_job jhp jhp

subgoal 2 (ID 2878) is:
 False
subgoal 3 (ID 2744) is:
 False

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


              by apply (H_priority_is_reflexive 0).

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

2 subgoals (ID 2878)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  EQUALt1 : t1 == predn t2
  SCHEDj : job_scheduled_at j t1 = false
  ============================
  False

subgoal 2 (ID 2744) is:
 False

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


          + eapply NOTHP, (H_respects_policy _ _ t2.-1); auto;
              last by rewrite -(eqbool_to_eqprop EQUALt1).

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

2 subgoals (ID 3481)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  EQUALt1 : t1 == predn t2
  SCHEDj : job_scheduled_at j t1 = false
  ============================
  backlogged sched j (predn t2)

subgoal 2 (ID 2744) is:
 False

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


            apply/andP; split; last first.

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

3 subgoals (ID 3545)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  EQUALt1 : t1 == predn t2
  SCHEDj : job_scheduled_at j t1 = false
  ============================
  ~~ scheduled_at sched j (predn t2)

subgoal 2 (ID 3544) is:
 job_ready sched j (predn t2)
subgoal 3 (ID 2744) is:
 False

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


            × by rewrite -(eqbool_to_eqprop EQUALt1); unfold job_scheduled_at in *; rewrite SCHEDj.

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

2 subgoals (ID 3544)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  EQUALt1 : t1 == predn t2
  SCHEDj : job_scheduled_at j t1 = false
  ============================
  job_ready sched j (predn t2)

subgoal 2 (ID 2744) is:
 False

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


            × have EQ: t1 = job_arrival j.

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

3 subgoals (ID 3562)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  EQUALt1 : t1 == predn t2
  SCHEDj : job_scheduled_at j t1 = false
  ============================
  t1 = job_arrival j

subgoal 2 (ID 3564) is:
 job_ready sched j (predn t2)
subgoal 3 (ID 2744) is:
 False

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


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

1 subgoal (ID 3562)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  EQUALt1 : t1 == predn t2
  SCHEDj : job_scheduled_at j t1 = false
  ============================
  t1 = job_arrival j

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


rewrite -eqSS in EQUALt1.

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

1 subgoal (ID 3614)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  SCHEDj : job_scheduled_at j t1 = false
  EQUALt1 : succn t1 == succn (predn t2)
  ============================
  t1 = job_arrival j

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


                have EQ: t2 = t1.+1.

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

2 subgoals (ID 3617)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  SCHEDj : job_scheduled_at j t1 = false
  EQUALt1 : succn t1 == succn (predn t2)
  ============================
  t2 = succn t1

subgoal 2 (ID 3619) is:
 t1 = job_arrival j

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


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

1 subgoal (ID 3617)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  SCHEDj : job_scheduled_at j t1 = false
  EQUALt1 : succn t1 == succn (predn t2)
  ============================
  t2 = succn t1

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


rewrite prednK in EQUALt1; first by apply/eqP; rewrite eq_sym.

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

1 subgoal (ID 3711)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  SCHEDj : job_scheduled_at j t1 = false
  ============================
  0 < t2

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


                  apply negbNE; rewrite -eqn0Ngt; apply/neqP; intros EQ0.

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

1 subgoal (ID 3802)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  SCHEDj : job_scheduled_at j t1 = false
  EQ0 : t2 = 0
  ============================
  False

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


                  move: INBI; rewrite EQ0; move ⇒ /andP [_ CONTR].

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

1 subgoal (ID 3846)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  SCHEDj : job_scheduled_at j t1 = false
  EQ0 : t2 = 0
  CONTR : job_arrival j < 0
  ============================
  False

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


                    by rewrite ltn0 in CONTR.

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

3 subgoals (ID 3619)

subgoal 1 (ID 3619) is:
 t1 = job_arrival j
subgoal 2 (ID 3564) is:
 job_ready sched j (predn t2)
subgoal 3 (ID 2744) is:
 False

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


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

1 subgoal (ID 3619)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  SCHEDj : job_scheduled_at j t1 = false
  EQUALt1 : succn t1 == succn (predn t2)
  EQ : t2 = succn t1
  ============================
  t1 = job_arrival j

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


clear EQUALt1.

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

1 subgoal (ID 3923)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  SCHEDj : job_scheduled_at j t1 = false
  EQ : t2 = succn t1
  ============================
  t1 = job_arrival j

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


                  by move: INBI; rewrite EQ ltnS -eqn_leq; move ⇒ /eqP INBI.

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

2 subgoals (ID 3564)

subgoal 1 (ID 3564) is:
 job_ready sched j (predn t2)
subgoal 2 (ID 2744) is:
 False

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


              }

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

2 subgoals (ID 3564)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  jhp : Job
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  Sched_jhp : scheduled_at sched jhp t1
  EQUALt1 : t1 == predn t2
  SCHEDj : job_scheduled_at j t1 = false
  EQ : t1 = job_arrival j
  ============================
  job_ready sched j (predn t2)

subgoal 2 (ID 2744) is:
 False

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


                by rewrite -(eqbool_to_eqprop EQUALt1) EQ; eapply job_pending_at_arrival; eauto 2.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2744)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  LARGERt1 : t1 < t
  ============================
  False

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



        - feed (NOTQUIET t); first by apply/andP; split.

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

1 subgoal (ID 3997)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  LARGERt1 : t1 < t
  ============================
  False

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


          apply NOTQUIET; intros j_hp' IN HP ARR.

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

1 subgoal (ID 4030)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  LARGERt1 : t1 < t
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' t
  ============================
  completed_by sched j_hp' t

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


          apply contraT; move ⇒ /negP NOTCOMP'; exfalso.

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

1 subgoal (ID 4058)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  LARGERt1 : t1 < t
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' t
  NOTCOMP' : ~ completed_by sched j_hp' t
  ============================
  False

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


          have BACK: backlogged sched j_hp' t.

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

2 subgoals (ID 4065)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  LARGERt1 : t1 < t
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' t
  NOTCOMP' : ~ completed_by sched j_hp' t
  ============================
  backlogged sched j_hp' t

subgoal 2 (ID 4067) is:
 False

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


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

1 subgoal (ID 4065)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  LARGERt1 : t1 < t
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' t
  NOTCOMP' : ~ completed_by sched j_hp' t
  ============================
  backlogged sched j_hp' t

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


apply/andP; split.

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

2 subgoals (ID 4093)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  LARGERt1 : t1 < t
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' t
  NOTCOMP' : ~ completed_by sched j_hp' t
  ============================
  job_ready sched j_hp' t

subgoal 2 (ID 4094) is:
 ~~ scheduled_at sched j_hp' t

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


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

3 subgoals (ID 4120)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  LARGERt1 : t1 < t
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' t
  NOTCOMP' : ~ completed_by sched j_hp' t
  ============================
  has_arrived j_hp' t

subgoal 2 (ID 4121) is:
 ~~ completed_by sched j_hp' t
subgoal 3 (ID 4094) is:
 ~~ scheduled_at sched j_hp' t

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


unfold arrived_before, has_arrived in ×.
(* ----------------------------------[ coqtop ]---------------------------------

3 subgoals (ID 4123)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  LARGERt1 : t1 < t
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : job_arrival j_hp' < t
  NOTCOMP' : ~ completed_by sched j_hp' t
  ============================
  job_arrival j_hp' <= t

subgoal 2 (ID 4121) is:
 ~~ completed_by sched j_hp' t
subgoal 3 (ID 4094) is:
 ~~ scheduled_at sched j_hp' t

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


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

2 subgoals (ID 4121)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  LARGERt1 : t1 < t
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' t
  NOTCOMP' : ~ completed_by sched j_hp' t
  ============================
  ~~ completed_by sched j_hp' t

subgoal 2 (ID 4094) is:
 ~~ scheduled_at sched j_hp' t

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


              apply/negP; intro COMP; apply NOTCOMP'.

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

2 subgoals (ID 4152)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  LARGERt1 : t1 < t
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' t
  NOTCOMP' : ~ completed_by sched j_hp' t
  COMP : completed_by sched j_hp' t
  ============================
  completed_by sched j_hp' t

subgoal 2 (ID 4094) is:
 ~~ scheduled_at sched j_hp' t

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


                by apply completion_monotonic with (t0 := t).

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

1 subgoal (ID 4094)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  LARGERt1 : t1 < t
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' t
  NOTCOMP' : ~ completed_by sched j_hp' t
  ============================
  ~~ scheduled_at sched j_hp' t

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


            - apply/negP; intro SCHED'.

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

1 subgoal (ID 4180)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  LARGERt1 : t1 < t
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' t
  NOTCOMP' : ~ completed_by sched j_hp' t
  SCHED' : scheduled_at sched j_hp' t
  ============================
  False

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


              move: (ideal_proc_model_is_a_uniprocessor_model jhp j_hp' sched t Sched_jhp SCHED') ⇒ EQ; subst.

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

1 subgoal (ID 4213)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_t_preemption_time : preemption_time sched (predn t2)
  H_t_in_busy_interval : t1 <= predn t2 < t2
  LEt : predn t2 < t2
  j_hp' : Job
  Sched_jhp : scheduled_at sched j_hp' (predn t2)
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : ~ quiet_time arr_seq sched j (predn t2)
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job j_hp' j
  LARGERt1 : t1 < predn t2
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  SCHED' : scheduled_at sched j_hp' (predn t2)
  NOTCOMP' : ~ completed_by sched j_hp' (predn t2)
  ARR : arrived_before j_hp' (predn t2)
  ============================
  False

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


                by apply: NOTHP.

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

1 subgoal (ID 4067)

subgoal 1 (ID 4067) is:
 False

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


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

1 subgoal (ID 4067)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  EQUALt2m1 : t = predn t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  NOTHP : ~ hep_job jhp j
  LARGERt1 : t1 < t
  j_hp' : Job
  IN : arrives_in arr_seq j_hp'
  HP : hep_job j_hp' j
  ARR : arrived_before j_hp' t
  NOTCOMP' : ~ completed_by sched j_hp' t
  BACK : backlogged sched j_hp' t
  ============================
  False

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



          apply NOTHP, (H_priority_is_transitive t j_hp'); eauto 2.

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

No more subgoals.

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


      Qed.

By combining the above facts we conclude that a job that is scheduled at time [t] has higher-or-equal priority.
      Corollary scheduled_at_preemption_time_implies_higher_or_equal_priority:
         jhp,
          scheduled_at sched jhp t
          hep_job jhp j.

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

1 subgoal (ID 2519)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  ============================
  forall jhp : Job, scheduled_at sched jhp t -> hep_job jhp j

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


      Proof.
        move: (H_t_in_busy_interval) (H_busy_interval_prefix) ⇒ /andP [GEt LEt] [SL [QUIET [NOTQUIET INBI]]].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2592)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  ============================
  forall jhp : Job, scheduled_at sched jhp t -> hep_job jhp j

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


        destruct t_lt_t2_or_t_eq_t2.

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

2 subgoals (ID 2599)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  H6 : t < predn t2
  ============================
  forall jhp : Job, scheduled_at sched jhp t -> hep_job jhp j

subgoal 2 (ID 2600) is:
 forall jhp : Job, scheduled_at sched jhp t -> hep_job jhp j

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


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

1 subgoal (ID 2600)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  GEt : t1 <= t
  LEt : t < t2
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  H6 : t = predn t2
  ============================
  forall jhp : Job, scheduled_at sched jhp t -> hep_job jhp j

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


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

No more subgoals.

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


      Qed.

Since a job that is scheduled at time [t] has higher-or-equal priority, by properties of a busy interval it cannot arrive before time instant [t1].
      Lemma scheduled_at_preemption_time_implies_arrived_between_within_busy_interval:
         jhp,
          scheduled_at sched jhp t
          arrived_between jhp t1 t2.

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

1 subgoal (ID 2528)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  ============================
  forall jhp : Job, scheduled_at sched jhp t -> arrived_between jhp t1 t2

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


      Proof.
        intros jhp Sched_jhp.

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

1 subgoal (ID 2530)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  ============================
  arrived_between jhp t1 t2

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


        rename H_work_conserving into WORK, H_jobs_come_from_arrival_sequence into CONS.

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

1 subgoal (ID 2531)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  WORK : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  ============================
  arrived_between jhp t1 t2

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


        move: (H_busy_interval_prefix) ⇒ [SL [QUIET [NOTQUIET INBI]]].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2563)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  WORK : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  ============================
  arrived_between jhp t1 t2

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


        move: (H_t_in_busy_interval) ⇒ /andP [GEt LEt].

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

1 subgoal (ID 2604)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  WORK : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  ============================
  arrived_between jhp t1 t2

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


        have HP := scheduled_at_preemption_time_implies_higher_or_equal_priority _ Sched_jhp.

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

1 subgoal (ID 2611)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  WORK : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  HP : hep_job jhp j
  ============================
  arrived_between jhp t1 t2

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


        move: (Sched_jhp) ⇒ PENDING.

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

1 subgoal (ID 2613)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  WORK : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  HP : hep_job jhp j
  PENDING : scheduled_at sched jhp t
  ============================
  arrived_between jhp t1 t2

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


        eapply scheduled_implies_pending in PENDING; eauto 2 with basic_facts.

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

1 subgoal (ID 2620)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  WORK : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  HP : hep_job jhp j
  PENDING : pending sched jhp t
  ============================
  arrived_between jhp t1 t2

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


        apply/andP; split; last by apply leq_ltn_trans with (n := t); first by move: PENDING ⇒ /andP [ARR _].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2657)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  WORK : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  HP : hep_job jhp j
  PENDING : pending sched jhp t
  ============================
  t1 <= job_arrival jhp

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


        apply contraT; rewrite -ltnNge; intro LT; exfalso.

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

1 subgoal (ID 2709)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  WORK : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  HP : hep_job jhp j
  PENDING : pending sched jhp t
  LT : job_arrival jhp < t1
  ============================
  False

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


        feed (QUIET jhp); first by eapply CONS, Sched_jhp.

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

1 subgoal (ID 2715)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  WORK : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  SL : t1 < t2
  QUIET : hep_job jhp j -> arrived_before jhp t1 -> completed_by sched jhp t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  HP : hep_job jhp j
  PENDING : pending sched jhp t
  LT : job_arrival jhp < t1
  ============================
  False

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


        specialize (QUIET HP LT).

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

1 subgoal (ID 2719)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  WORK : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  SL : t1 < t2
  QUIET : completed_by sched jhp t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  HP : hep_job jhp j
  PENDING : pending sched jhp t
  LT : job_arrival jhp < t1
  ============================
  False

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


        have COMP: job_completed_by jhp t by apply completion_monotonic with (t0 := t1).

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

1 subgoal (ID 2728)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  WORK : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  SL : t1 < t2
  QUIET : completed_by sched jhp t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  HP : hep_job jhp j
  PENDING : pending sched jhp t
  LT : job_arrival jhp < t1
  COMP : job_completed_by jhp t
  ============================
  False

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


        apply completed_implies_not_scheduled in COMP; last by done.

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

1 subgoal (ID 2733)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  WORK : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  SL : t1 < t2
  QUIET : completed_by sched jhp t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  HP : hep_job jhp j
  PENDING : pending sched jhp t
  LT : job_arrival jhp < t1
  COMP : ~~ scheduled_at sched jhp t
  ============================
  False

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


          by move: COMP ⇒ /negP COMP; apply COMP.

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

No more subgoals.

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


      Qed.

From the above lemmas we prove that there is a job [jhp] that is (1) scheduled at time [t], (2) has higher-or-equal priority, and (3) arrived between [t1] and [t2].
      Corollary not_quiet_implies_exists_scheduled_hp_job_at_preemption_point:
         j_hp,
          arrived_between j_hp t1 t2
          hep_job j_hp j
          job_scheduled_at j_hp t.

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

1 subgoal (ID 2537)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  ============================
  exists j_hp : Job,
    arrived_between j_hp t1 t2 /\ hep_job j_hp j /\ job_scheduled_at j_hp t

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


      Proof.
        move: (H_busy_interval_prefix) ⇒ [SL [QUIET [NOTQUIET INBI]]].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2569)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  ============================
  exists j_hp : Job,
    arrived_between j_hp t1 t2 /\ hep_job j_hp j /\ job_scheduled_at j_hp t

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


        move: (H_t_in_busy_interval) ⇒ /andP [GEt LEt].

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

1 subgoal (ID 2610)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  ============================
  exists j_hp : Job,
    arrived_between j_hp t1 t2 /\ hep_job j_hp j /\ job_scheduled_at j_hp t

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


        ideal_proc_model_sched_case_analysis sched t jhp.

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

2 subgoals (ID 2618)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  Idle : is_idle sched t
  ============================
  exists j_hp : Job,
    arrived_between j_hp t1 t2 /\ hep_job j_hp j /\ job_scheduled_at j_hp t

subgoal 2 (ID 2623) is:
 exists j_hp : Job,
   arrived_between j_hp t1 t2 /\ hep_job j_hp j /\ job_scheduled_at j_hp t

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


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

1 subgoal (ID 2618)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  Idle : is_idle sched t
  ============================
  exists j_hp : Job,
    arrived_between j_hp t1 t2 /\ hep_job j_hp j /\ job_scheduled_at j_hp t

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


by exfalso; apply instant_t_is_not_idle.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2623)

subgoal 1 (ID 2623) is:
 exists j_hp : Job,
   arrived_between j_hp t1 t2 /\ hep_job j_hp j /\ job_scheduled_at j_hp t

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


}

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

1 subgoal (ID 2623)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  ============================
  exists j_hp : Job,
    arrived_between j_hp t1 t2 /\ hep_job j_hp j /\ job_scheduled_at j_hp t

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


         jhp.

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

1 subgoal (ID 2627)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  ============================
  arrived_between jhp t1 t2 /\ hep_job jhp j /\ job_scheduled_at jhp t

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


        repeat split.

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

3 subgoals (ID 2629)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  ============================
  arrived_between jhp t1 t2

subgoal 2 (ID 2633) is:
 hep_job jhp j
subgoal 3 (ID 2634) is:
 job_scheduled_at jhp t

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


        - by apply scheduled_at_preemption_time_implies_arrived_between_within_busy_interval.

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

2 subgoals (ID 2633)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  ============================
  hep_job jhp j

subgoal 2 (ID 2634) is:
 job_scheduled_at jhp t

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


        - by apply scheduled_at_preemption_time_implies_higher_or_equal_priority.

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

1 subgoal (ID 2634)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  t : instant
  H_t_in_busy_interval : t1 <= t < t2
  H_t_preemption_time : preemption_time sched t
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  INBI : t1 <= job_arrival j < t2
  GEt : t1 <= t
  LEt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  ============================
  job_scheduled_at jhp t

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


        - by done.

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

No more subgoals.

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


      Qed.

    End ProcessorBusyWithHEPJobAtPreemptionPoints.

Next we prove that every nonpreemptive segment always begins with a preemption time ...
    Lemma scheduling_of_any_segment_starts_with_preemption_time:
       j t,
        job_scheduled_at j t
         pt,
          job_arrival j pt t
          preemption_time sched pt
          ( t', pt t' t job_scheduled_at j t').

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

1 subgoal (ID 2493)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  ============================
  forall (j0 : Job) (t : instant),
  job_scheduled_at j0 t ->
  exists pt : nat,
    job_arrival j0 <= pt <= t /\
    preemption_time sched pt /\
    (forall t' : nat, pt <= t' <= t -> job_scheduled_at j0 t')

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


    Proof.
      intros s t SCHEDst.

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

1 subgoal (ID 2496)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  ============================
  exists pt : nat,
    job_arrival s <= pt <= t /\
    preemption_time sched pt /\
    (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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


      have EX: t', (t' t) && (job_scheduled_at s t')
                              && (all (fun t''job_scheduled_at s t'') (index_iota t' t.+1 )).

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

2 subgoals (ID 2501)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  ============================
  exists t' : nat,
    (t' <= t) && job_scheduled_at s t' &&
    all [eta job_scheduled_at s] (index_iota t' (succn t))

subgoal 2 (ID 2503) is:
 exists pt : nat,
   job_arrival s <= pt <= t /\
   preemption_time sched pt /\
   (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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


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

1 subgoal (ID 2501)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  ============================
  exists t' : nat,
    (t' <= t) && job_scheduled_at s t' &&
    all [eta job_scheduled_at s] (index_iota t' (succn t))

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


t.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2505)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  ============================
  (t <= t) && job_scheduled_at s t &&
  all [eta job_scheduled_at s] (index_iota t (succn t))

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


apply/andP; split; [ by apply/andP; split | ].

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

1 subgoal (ID 2532)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  ============================
  all [eta job_scheduled_at s] (index_iota t (succn t))

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


        apply/allP; intros t'.

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

1 subgoal (ID 2589)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  t' : nat_eqType
  ============================
  t' \in index_iota t (succn t) -> job_scheduled_at s t'

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


          by rewrite mem_index_iota ltnS -eqn_leq; move ⇒ /eqP <-.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2503)

subgoal 1 (ID 2503) is:
 exists pt : nat,
   job_arrival s <= pt <= t /\
   preemption_time sched pt /\
   (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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



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

1 subgoal (ID 2503)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  EX : exists t' : nat,
         (t' <= t) && job_scheduled_at s t' &&
         all [eta job_scheduled_at s] (index_iota t' (succn t))
  ============================
  exists pt : nat,
    job_arrival s <= pt <= t /\
    preemption_time sched pt /\
    (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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



      have MIN := ex_minnP EX.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2645)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  EX : exists t' : nat,
         (t' <= t) && job_scheduled_at s t' &&
         all [eta job_scheduled_at s] (index_iota t' (succn t))
  MIN : ex_minn_spec
          (fun t' : nat =>
           (t' <= t) && job_scheduled_at s t' &&
           all [eta job_scheduled_at s] (index_iota t' (succn t)))
          (ex_minn
             (P:=fun t' : nat =>
                 (t' <= t) && job_scheduled_at s t' &&
                 all [eta job_scheduled_at s] (index_iota t' (succn t))) EX)
  ============================
  exists pt : nat,
    job_arrival s <= pt <= t /\
    preemption_time sched pt /\
    (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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


      move: MIN ⇒ [mpt /andP [/andP [LT1 SCHEDsmpt] /allP ALL] MIN]; clear EX.

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

1 subgoal (ID 2769)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt <= t
  SCHEDsmpt : job_scheduled_at s mpt
  ALL : {in index_iota mpt (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt <= n
  ============================
  exists pt : nat,
    job_arrival s <= pt <= t /\
    preemption_time sched pt /\
    (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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


      destruct mpt.

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

2 subgoals (ID 2787)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  LT1 : 0 <= t
  SCHEDsmpt : job_scheduled_at s 0
  ALL : {in index_iota 0 (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        0 <= n
  ============================
  exists pt : nat,
    job_arrival s <= pt <= t /\
    preemption_time sched pt /\
    (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

subgoal 2 (ID 2792) is:
 exists pt : nat,
   job_arrival s <= pt <= t /\
   preemption_time sched pt /\
   (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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


      - 0; repeat split.

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

4 subgoals (ID 2796)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  LT1 : 0 <= t
  SCHEDsmpt : job_scheduled_at s 0
  ALL : {in index_iota 0 (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        0 <= n
  ============================
  job_arrival s <= 0 <= t

subgoal 2 (ID 2800) is:
 preemption_time sched 0
subgoal 3 (ID 2801) is:
 forall t' : nat, 0 <= t' <= t -> job_scheduled_at s t'
subgoal 4 (ID 2792) is:
 exists pt : nat,
   job_arrival s <= pt <= t /\
   preemption_time sched pt /\
   (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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


        + by apply/andP; split; first apply H_jobs_must_arrive_to_execute in SCHEDsmpt.

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

3 subgoals (ID 2800)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  LT1 : 0 <= t
  SCHEDsmpt : job_scheduled_at s 0
  ALL : {in index_iota 0 (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        0 <= n
  ============================
  preemption_time sched 0

subgoal 2 (ID 2801) is:
 forall t' : nat, 0 <= t' <= t -> job_scheduled_at s t'
subgoal 3 (ID 2792) is:
 exists pt : nat,
   job_arrival s <= pt <= t /\
   preemption_time sched pt /\
   (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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


        + by eapply zero_is_pt; eauto 2.

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

2 subgoals (ID 2801)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  LT1 : 0 <= t
  SCHEDsmpt : job_scheduled_at s 0
  ALL : {in index_iota 0 (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        0 <= n
  ============================
  forall t' : nat, 0 <= t' <= t -> job_scheduled_at s t'

subgoal 2 (ID 2792) is:
 exists pt : nat,
   job_arrival s <= pt <= t /\
   preemption_time sched pt /\
   (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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


        + by intros; apply ALL; rewrite mem_iota subn0 add0n ltnS.

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

1 subgoal (ID 2792)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  ============================
  exists pt : nat,
    job_arrival s <= pt <= t /\
    preemption_time sched pt /\
    (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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


      - have NSCHED: ~~ job_scheduled_at s mpt.

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

2 subgoals (ID 2866)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  ============================
  ~~ job_scheduled_at s mpt

subgoal 2 (ID 2868) is:
 exists pt : nat,
   job_arrival s <= pt <= t /\
   preemption_time sched pt /\
   (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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


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

1 subgoal (ID 2866)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  ============================
  ~~ job_scheduled_at s mpt

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


apply/negP; intros SCHED.

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

1 subgoal (ID 2890)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  SCHED : job_scheduled_at s mpt
  ============================
  False

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


          enough (F : mpt < mpt); first by rewrite ltnn in F.

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

1 subgoal (ID 2892)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  SCHED : job_scheduled_at s mpt
  ============================
  mpt < mpt

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


          apply MIN.

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

1 subgoal (ID 2969)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  SCHED : job_scheduled_at s mpt
  ============================
  (mpt <= t) && job_scheduled_at s mpt &&
  all [eta job_scheduled_at s] (index_iota mpt (succn t))

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


          apply/andP; split; [by apply/andP; split; [ apply ltnW | ] | ].

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

1 subgoal (ID 2996)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  SCHED : job_scheduled_at s mpt
  ============================
  all [eta job_scheduled_at s] (index_iota mpt (succn t))

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


          apply/allP; intros t'.

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

1 subgoal (ID 3054)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  SCHED : job_scheduled_at s mpt
  t' : nat_eqType
  ============================
  t' \in index_iota mpt (succn t) -> job_scheduled_at s t'

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


          rewrite mem_index_iota; move ⇒ /andP [GE LE].

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

1 subgoal (ID 3099)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  SCHED : job_scheduled_at s mpt
  t' : nat_eqType
  GE : mpt <= t'
  LE : t' < succn t
  ============================
  job_scheduled_at s t'

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


          move: GE; rewrite leq_eqVlt; move ⇒ /orP [/eqP EQ| LT].

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

2 subgoals (ID 3180)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  SCHED : job_scheduled_at s mpt
  t' : nat_eqType
  LE : t' < succn t
  EQ : mpt = t'
  ============================
  job_scheduled_at s t'

subgoal 2 (ID 3181) is:
 job_scheduled_at s t'

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


          - by subst t'.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 3181)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  SCHED : job_scheduled_at s mpt
  t' : nat_eqType
  LE : t' < succn t
  LT : mpt < t'
  ============================
  job_scheduled_at s t'

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


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

1 subgoal (ID 2868)

subgoal 1 (ID 2868) is:
 exists pt : nat,
   job_arrival s <= pt <= t /\
   preemption_time sched pt /\
   (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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


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

1 subgoal (ID 2868)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  NSCHED : ~~ job_scheduled_at s mpt
  ============================
  exists pt : nat,
    job_arrival s <= pt <= t /\
    preemption_time sched pt /\
    (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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


        
        have PP: preemption_time sched mpt.+1.

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

2 subgoals (ID 3223)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  NSCHED : ~~ job_scheduled_at s mpt
  ============================
  preemption_time sched (succn mpt)

subgoal 2 (ID 3225) is:
 exists pt : nat,
   job_arrival s <= pt <= t /\
   preemption_time sched pt /\
   (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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


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

1 subgoal (ID 3223)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  NSCHED : ~~ job_scheduled_at s mpt
  ============================
  preemption_time sched (succn mpt)

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


by eapply first_moment_is_pt with (j0 := s); eauto 2.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 3225)

subgoal 1 (ID 3225) is:
 exists pt : nat,
   job_arrival s <= pt <= t /\
   preemption_time sched pt /\
   (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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


}

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

1 subgoal (ID 3225)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  NSCHED : ~~ job_scheduled_at s mpt
  PP : preemption_time sched (succn mpt)
  ============================
  exists pt : nat,
    job_arrival s <= pt <= t /\
    preemption_time sched pt /\
    (forall t' : nat, pt <= t' <= t -> job_scheduled_at s t')

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


         mpt.+1; repeat split; try done.

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

2 subgoals (ID 3250)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  NSCHED : ~~ job_scheduled_at s mpt
  PP : preemption_time sched (succn mpt)
  ============================
  job_arrival s <= succn mpt <= t

subgoal 2 (ID 3255) is:
 forall t' : nat, mpt < t' <= t -> job_scheduled_at s t'

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


        + apply/andP; split; last by done.

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

2 subgoals (ID 3333)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  NSCHED : ~~ job_scheduled_at s mpt
  PP : preemption_time sched (succn mpt)
  ============================
  job_arrival s <= succn mpt

subgoal 2 (ID 3255) is:
 forall t' : nat, mpt < t' <= t -> job_scheduled_at s t'

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


            by apply H_jobs_must_arrive_to_execute in SCHEDsmpt.

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

1 subgoal (ID 3255)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  NSCHED : ~~ job_scheduled_at s mpt
  PP : preemption_time sched (succn mpt)
  ============================
  forall t' : nat, mpt < t' <= t -> job_scheduled_at s t'

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


        + movet' /andP [GE LE].

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

1 subgoal (ID 3377)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  s : Job
  t : instant
  SCHEDst : job_scheduled_at s t
  mpt : nat
  LT1 : mpt < t
  SCHEDsmpt : job_scheduled_at s (succn mpt)
  ALL : {in index_iota (succn mpt) (succn t),
          forall x : nat_eqType, job_scheduled_at s x}
  MIN : forall n : nat,
        (n <= t) && job_scheduled_at s n &&
        all [eta job_scheduled_at s] (index_iota n (succn t)) -> 
        mpt < n
  NSCHED : ~~ job_scheduled_at s mpt
  PP : preemption_time sched (succn mpt)
  t' : nat
  GE : mpt < t'
  LE : t' <= t
  ============================
  job_scheduled_at s t'

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


            by apply ALL; rewrite mem_index_iota; apply/andP; split.

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

No more subgoals.

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


    Qed.

... and the fact that at any time instant after a preemption point the processor is always busy with a job with higher or equal priority.
    Lemma not_quiet_implies_exists_scheduled_hp_job_after_preemption_point:
       tp t,
        preemption_time sched tp
        t1 tp < t2
        tp t < t2
         j_hp,
          arrived_between j_hp t1 t.+1
          hep_job j_hp j
          job_scheduled_at j_hp t.

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

1 subgoal (ID 2506)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  ============================
  forall (tp : instant) (t : nat),
  preemption_time sched tp ->
  t1 <= tp < t2 ->
  tp <= t < t2 ->
  exists j_hp : Job,
    arrived_between j_hp t1 (succn t) /\
    hep_job j_hp j /\ job_scheduled_at j_hp t

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


    Proof.
      move: (H_jobs_come_from_arrival_sequence) (H_work_conserving) ⇒ CONS WORK.

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

1 subgoal (ID 2510)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  ============================
  forall (tp : instant) (t : nat),
  preemption_time sched tp ->
  t1 <= tp < t2 ->
  tp <= t < t2 ->
  exists j_hp : Job,
    arrived_between j_hp t1 (succn t) /\
    hep_job j_hp j /\ job_scheduled_at j_hp t

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


      move: (H_respects_policy) ⇒ PRIO.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2512)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  ============================
  forall (tp : instant) (t : nat),
  preemption_time sched tp ->
  t1 <= tp < t2 ->
  tp <= t < t2 ->
  exists j_hp : Job,
    arrived_between j_hp t1 (succn t) /\
    hep_job j_hp j /\ job_scheduled_at j_hp t

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



      movetp t PRPOINT /andP [GEtp LTtp] /andP [LEtp LTt].

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

1 subgoal (ID 2594)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  ============================
  exists j_hp : Job,
    arrived_between j_hp t1 (succn t) /\
    hep_job j_hp j /\ job_scheduled_at j_hp t

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


      ideal_proc_model_sched_case_analysis_eq sched t jhp.

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

2 subgoals (ID 2642)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  Idle : is_idle sched t
  EqIdle : sched t = None
  ============================
  exists j_hp : Job,
    arrived_between j_hp t1 (succn t) /\
    hep_job j_hp j /\ job_scheduled_at j_hp t

subgoal 2 (ID 2688) is:
 exists j_hp : Job,
   arrived_between j_hp t1 (succn t) /\
   hep_job j_hp j /\ job_scheduled_at j_hp t

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


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

1 subgoal (ID 2642)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  Idle : is_idle sched t
  EqIdle : sched t = None
  ============================
  exists j_hp : Job,
    arrived_between j_hp t1 (succn t) /\
    hep_job j_hp j /\ job_scheduled_at j_hp t

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


exfalso; eapply not_quiet_implies_not_idle with (t0 := t); eauto 2.

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

1 subgoal (ID 2707)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  Idle : is_idle sched t
  EqIdle : sched t = None
  ============================
  t1 <= t < t2

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


          by apply/andP; split; first apply leq_trans with tp.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2688)

subgoal 1 (ID 2688) is:
 exists j_hp : Job,
   arrived_between j_hp t1 (succn t) /\
   hep_job j_hp j /\ job_scheduled_at j_hp t

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


}

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

1 subgoal (ID 2688)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  ============================
  exists j_hp : Job,
    arrived_between j_hp t1 (succn t) /\
    hep_job j_hp j /\ job_scheduled_at j_hp t

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


       jhp.

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

1 subgoal (ID 2754)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  ============================
  arrived_between jhp t1 (succn t) /\ hep_job jhp j /\ job_scheduled_at jhp t

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


      have HP: hep_job jhp j.

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

2 subgoals (ID 2757)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  ============================
  hep_job jhp j

subgoal 2 (ID 2759) is:
 arrived_between jhp t1 (succn t) /\ hep_job jhp j /\ job_scheduled_at jhp t

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


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

1 subgoal (ID 2757)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  ============================
  hep_job jhp j

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


intros.
        have SOAS := scheduling_of_any_segment_starts_with_preemption_time _ _ Sched_jhp.

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

1 subgoal (ID 2768)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  SOAS : exists pt : nat,
           job_arrival jhp <= pt <= t /\
           preemption_time sched pt /\
           (forall t' : nat, pt <= t' <= t -> job_scheduled_at jhp t')
  ============================
  hep_job jhp j

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


        move: SOAS ⇒ [prt [/andP [_ LE] [PR SCH]]].

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

1 subgoal (ID 2840)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  prt : nat
  LE : prt <= t
  PR : preemption_time sched prt
  SCH : forall t' : nat, prt <= t' <= t -> job_scheduled_at jhp t'
  ============================
  hep_job jhp j

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


        case E:(t1 prt).

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

2 subgoals (ID 2895)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  prt : nat
  LE : prt <= t
  PR : preemption_time sched prt
  SCH : forall t' : nat, prt <= t' <= t -> job_scheduled_at jhp t'
  E : (t1 <= prt) = true
  ============================
  hep_job jhp j

subgoal 2 (ID 2942) is:
 hep_job jhp j

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


        - move: E ⇒ /eqP /eqP E; rewrite subn_eq0 in E.

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

2 subgoals (ID 3086)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  prt : nat
  LE : prt <= t
  PR : preemption_time sched prt
  SCH : forall t' : nat, prt <= t' <= t -> job_scheduled_at jhp t'
  E : t1 <= prt
  ============================
  hep_job jhp j

subgoal 2 (ID 2942) is:
 hep_job jhp j

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


          edestruct not_quiet_implies_exists_scheduled_hp_job_at_preemption_point as [jlp [_ [HEP SCHEDjhp]]]; eauto 2.

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

3 subgoals (ID 3091)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  prt : nat
  LE : prt <= t
  PR : preemption_time sched prt
  SCH : forall t' : nat, prt <= t' <= t -> job_scheduled_at jhp t'
  E : t1 <= prt
  ============================
  t1 <= prt < t2

subgoal 2 (ID 3107) is:
 hep_job jhp j
subgoal 3 (ID 2942) is:
 hep_job jhp j

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


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

1 subgoal (ID 3091)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  prt : nat
  LE : prt <= t
  PR : preemption_time sched prt
  SCH : forall t' : nat, prt <= t' <= t -> job_scheduled_at jhp t'
  E : t1 <= prt
  ============================
  t1 <= prt < t2

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


by apply /andP; split; last by apply leq_ltn_trans with t.
(* ----------------------------------[ coqtop ]---------------------------------

3 subgoals (ID 3107)

subgoal 1 (ID 3107) is:
 hep_job jhp j
subgoal 2 (ID 2942) is:
 hep_job jhp j
subgoal 3 (ID 2759) is:
 arrived_between jhp t1 (succn t) /\ hep_job jhp j /\ job_scheduled_at jhp t

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


}

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

2 subgoals (ID 3107)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  prt : nat
  LE : prt <= t
  PR : preemption_time sched prt
  SCH : forall t' : nat, prt <= t' <= t -> job_scheduled_at jhp t'
  E : t1 <= prt
  jlp : Job
  HEP : hep_job jlp j
  SCHEDjhp : job_scheduled_at jlp prt
  ============================
  hep_job jhp j

subgoal 2 (ID 2942) is:
 hep_job jhp j

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


          enough (EQ : jhp = jlp); first by subst.

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

2 subgoals (ID 3156)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  prt : nat
  LE : prt <= t
  PR : preemption_time sched prt
  SCH : forall t' : nat, prt <= t' <= t -> job_scheduled_at jhp t'
  E : t1 <= prt
  jlp : Job
  HEP : hep_job jlp j
  SCHEDjhp : job_scheduled_at jlp prt
  ============================
  jhp = jlp

subgoal 2 (ID 2942) is:
 hep_job jhp j

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


          eapply ideal_proc_model_is_a_uniprocessor_model with (t0 := prt); eauto;
            by apply SCH; apply/andP; split.

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

1 subgoal (ID 2942)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  prt : nat
  LE : prt <= t
  PR : preemption_time sched prt
  SCH : forall t' : nat, prt <= t' <= t -> job_scheduled_at jhp t'
  E : (t1 <= prt) = false
  ============================
  hep_job jhp j

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


        - move: E ⇒ /eqP /neqP E; rewrite -lt0n subn_gt0 in E.

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

1 subgoal (ID 4279)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  prt : nat
  LE : prt <= t
  PR : preemption_time sched prt
  SCH : forall t' : nat, prt <= t' <= t -> job_scheduled_at jhp t'
  E : prt < t1
  ============================
  hep_job jhp j

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


          apply negbNE; apply/negP; intros LP; rename jhp into jlp.

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

1 subgoal (ID 4303)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jlp : Job
  Sched_jhp : scheduled_at sched jlp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jlp (sched t) x) x x in F]| <> 0
  prt : nat
  LE : prt <= t
  PR : preemption_time sched prt
  SCH : forall t' : nat, prt <= t' <= t -> job_scheduled_at jlp t'
  E : prt < t1
  LP : ~~ hep_job jlp j
  ============================
  False

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


          edestruct not_quiet_implies_exists_scheduled_hp_job_at_preemption_point
            as [jhp [_ [HEP SCHEDjhp]]]; [ | apply PRPOINT| ]; first by apply/andP; split.

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

1 subgoal (ID 4324)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jlp : Job
  Sched_jhp : scheduled_at sched jlp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jlp (sched t) x) x x in F]| <> 0
  prt : nat
  LE : prt <= t
  PR : preemption_time sched prt
  SCH : forall t' : nat, prt <= t' <= t -> job_scheduled_at jlp t'
  E : prt < t1
  LP : ~~ hep_job jlp j
  jhp : Job
  HEP : hep_job jhp j
  SCHEDjhp : job_scheduled_at jhp tp
  ============================
  False

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


          move: LP ⇒ /negP LP; apply: LP.

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

1 subgoal (ID 4385)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jlp : Job
  Sched_jhp : scheduled_at sched jlp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jlp (sched t) x) x x in F]| <> 0
  prt : nat
  LE : prt <= t
  PR : preemption_time sched prt
  SCH : forall t' : nat, prt <= t' <= t -> job_scheduled_at jlp t'
  E : prt < t1
  jhp : Job
  HEP : hep_job jhp j
  SCHEDjhp : job_scheduled_at jhp tp
  ============================
  hep_job jlp j

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


          enough (EQ : jhp = jlp); first by subst.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 4389)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jlp : Job
  Sched_jhp : scheduled_at sched jlp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jlp (sched t) x) x x in F]| <> 0
  prt : nat
  LE : prt <= t
  PR : preemption_time sched prt
  SCH : forall t' : nat, prt <= t' <= t -> job_scheduled_at jlp t'
  E : prt < t1
  jhp : Job
  HEP : hep_job jhp j
  SCHEDjhp : job_scheduled_at jhp tp
  ============================
  jhp = jlp

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


          eapply ideal_proc_model_is_a_uniprocessor_model with (j1 := jhp) (t0 := tp); eauto.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 4402)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jlp : Job
  Sched_jhp : scheduled_at sched jlp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jlp (sched t) x) x x in F]| <> 0
  prt : nat
  LE : prt <= t
  PR : preemption_time sched prt
  SCH : forall t' : nat, prt <= t' <= t -> job_scheduled_at jlp t'
  E : prt < t1
  jhp : Job
  HEP : hep_job jhp j
  SCHEDjhp : job_scheduled_at jhp tp
  ============================
  scheduled_at sched jlp tp

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


            by apply SCH; apply/andP; split; first apply leq_trans with t1; auto.

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

1 subgoal (ID 2759)

subgoal 1 (ID 2759) is:
 arrived_between jhp t1 (succn t) /\ hep_job jhp j /\ job_scheduled_at jhp t

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


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

1 subgoal (ID 2759)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  HP : hep_job jhp j
  ============================
  arrived_between jhp t1 (succn t) /\ hep_job jhp j /\ job_scheduled_at jhp t

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



      repeat split; try done.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 5371)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  HP : hep_job jhp j
  ============================
  arrived_between jhp t1 (succn t)

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


      move: (H_busy_interval_prefix) ⇒ [SL [QUIET [NOTQUIET EXj]]]; move: (Sched_jhp) ⇒ PENDING.

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

1 subgoal (ID 5435)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  HP : hep_job jhp j
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  EXj : t1 <= job_arrival j < t2
  PENDING : scheduled_at sched jhp t
  ============================
  arrived_between jhp t1 (succn t)

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


      eapply scheduled_implies_pending in PENDING; eauto with basic_facts.

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

1 subgoal (ID 5442)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  HP : hep_job jhp j
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  EXj : t1 <= job_arrival j < t2
  PENDING : pending sched jhp t
  ============================
  arrived_between jhp t1 (succn t)

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


      apply/andP; split; last by apply leq_ltn_trans with (n := t); first by move: PENDING ⇒ /andP [ARR _].

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

1 subgoal (ID 7151)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  HP : hep_job jhp j
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  EXj : t1 <= job_arrival j < t2
  PENDING : pending sched jhp t
  ============================
  t1 <= job_arrival jhp

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


      apply contraT; rewrite -ltnNge; intro LT; exfalso.

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

1 subgoal (ID 7203)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  HP : hep_job jhp j
  SL : t1 < t2
  QUIET : quiet_time arr_seq sched j t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  EXj : t1 <= job_arrival j < t2
  PENDING : pending sched jhp t
  LT : job_arrival jhp < t1
  ============================
  False

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


      feed (QUIET jhp); first by eapply CONS, Sched_jhp.

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

1 subgoal (ID 7209)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  HP : hep_job jhp j
  SL : t1 < t2
  QUIET : hep_job jhp j -> arrived_before jhp t1 -> completed_by sched jhp t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  EXj : t1 <= job_arrival j < t2
  PENDING : pending sched jhp t
  LT : job_arrival jhp < t1
  ============================
  False

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


      specialize (QUIET HP LT).

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

1 subgoal (ID 7213)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  HP : hep_job jhp j
  SL : t1 < t2
  QUIET : completed_by sched jhp t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  EXj : t1 <= job_arrival j < t2
  PENDING : pending sched jhp t
  LT : job_arrival jhp < t1
  ============================
  False

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


      have COMP: job_completed_by jhp t.

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

2 subgoals (ID 7214)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  HP : hep_job jhp j
  SL : t1 < t2
  QUIET : completed_by sched jhp t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  EXj : t1 <= job_arrival j < t2
  PENDING : pending sched jhp t
  LT : job_arrival jhp < t1
  ============================
  job_completed_by jhp t

subgoal 2 (ID 7216) is:
 False

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


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

1 subgoal (ID 7214)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  HP : hep_job jhp j
  SL : t1 < t2
  QUIET : completed_by sched jhp t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  EXj : t1 <= job_arrival j < t2
  PENDING : pending sched jhp t
  LT : job_arrival jhp < t1
  ============================
  job_completed_by jhp t

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


by apply completion_monotonic with (t0 := t1); [ apply leq_trans with tp | ].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 7216)

subgoal 1 (ID 7216) is:
 False

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


}

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

1 subgoal (ID 7216)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  HP : hep_job jhp j
  SL : t1 < t2
  QUIET : completed_by sched jhp t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  EXj : t1 <= job_arrival j < t2
  PENDING : pending sched jhp t
  LT : job_arrival jhp < t1
  COMP : job_completed_by jhp t
  ============================
  False

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


      apply completed_implies_not_scheduled in COMP; last by done.

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

1 subgoal (ID 7229)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  CONS : jobs_come_from_arrival_sequence sched arr_seq
  WORK : work_conserving arr_seq sched
  PRIO : respects_policy_at_preemption_point arr_seq sched
  tp : instant
  t : nat
  PRPOINT : preemption_time sched tp
  GEtp : t1 <= tp
  LTtp : tp < t2
  LEtp : tp <= t
  LTt : t < t2
  jhp : Job
  Sched_jhp : scheduled_at sched jhp t
  EqSched_jhp : #|[pred x |
                    let
                    'FiniteQuant.Quantified F :=
                     FiniteQuant.ex (T:=Core)
                       (, scheduled_on jhp (sched t) x) x x in F]| <> 0
  HP : hep_job jhp j
  SL : t1 < t2
  QUIET : completed_by sched jhp t1
  NOTQUIET : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  EXj : t1 <= job_arrival j < t2
  PENDING : pending sched jhp t
  LT : job_arrival jhp < t1
  COMP : ~~ scheduled_at sched jhp t
  ============================
  False

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


        by move: COMP ⇒ /negP COMP; apply COMP.

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

No more subgoals.

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


    Qed.

Now, suppose there exists some constant K that bounds the distance to a preemption time from the beginning of the busy interval.
    Variable K : duration.
    Hypothesis H_preemption_time_exists:
       pr_t, preemption_time sched pr_t t1 pr_t t1 + K.

Then we prove that the processor is always busy with a job with higher-or-equal priority after time instant [t1 + K].
    Lemma not_quiet_implies_exists_scheduled_hp_job:
       t,
        t1 + K t < t2
         j_hp,
          arrived_between j_hp t1 t.+1
          hep_job j_hp j
          job_scheduled_at j_hp t.

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

1 subgoal (ID 2520)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  K : duration
  H_preemption_time_exists : exists pr_t : instant,
                               preemption_time sched pr_t /\
                               t1 <= pr_t <= t1 + K
  ============================
  forall t : nat,
  t1 + K <= t < t2 ->
  exists j_hp : Job,
    arrived_between j_hp t1 (succn t) /\
    hep_job j_hp j /\ job_scheduled_at j_hp t

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


    Proof.
      movet /andP [GE LT].

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

1 subgoal (ID 2561)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  K : duration
  H_preemption_time_exists : exists pr_t : instant,
                               preemption_time sched pr_t /\
                               t1 <= pr_t <= t1 + K
  t : nat
  GE : t1 + K <= t
  LT : t < t2
  ============================
  exists j_hp : Job,
    arrived_between j_hp t1 (succn t) /\
    hep_job j_hp j /\ job_scheduled_at j_hp t

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


      move: H_preemption_time_exists ⇒ [prt [PR /andP [GEprt LEprt]]].

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

1 subgoal (ID 2622)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  K : duration
  H_preemption_time_exists : exists pr_t : instant,
                               preemption_time sched pr_t /\
                               t1 <= pr_t <= t1 + K
  t : nat
  GE : t1 + K <= t
  LT : t < t2
  prt : instant
  PR : preemption_time sched prt
  GEprt : t1 <= prt
  LEprt : prt <= t1 + K
  ============================
  exists j_hp : Job,
    arrived_between j_hp t1 (succn t) /\
    hep_job j_hp j /\ job_scheduled_at j_hp t

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


      apply not_quiet_implies_exists_scheduled_hp_job_after_preemption_point with (tp := prt); eauto 2.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2624)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  K : duration
  H_preemption_time_exists : exists pr_t : instant,
                               preemption_time sched pr_t /\
                               t1 <= pr_t <= t1 + K
  t : nat
  GE : t1 + K <= t
  LT : t < t2
  prt : instant
  PR : preemption_time sched prt
  GEprt : t1 <= prt
  LEprt : prt <= t1 + K
  ============================
  t1 <= prt < t2

subgoal 2 (ID 2625) is:
 prt <= t < t2

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


      - apply/andP; split; first by done.

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

2 subgoals (ID 2668)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  K : duration
  H_preemption_time_exists : exists pr_t : instant,
                               preemption_time sched pr_t /\
                               t1 <= pr_t <= t1 + K
  t : nat
  GE : t1 + K <= t
  LT : t < t2
  prt : instant
  PR : preemption_time sched prt
  GEprt : t1 <= prt
  LEprt : prt <= t1 + K
  ============================
  prt < t2

subgoal 2 (ID 2625) is:
 prt <= t < t2

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


         apply leq_ltn_trans with (t1 + K); first by done.

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

2 subgoals (ID 2670)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  K : duration
  H_preemption_time_exists : exists pr_t : instant,
                               preemption_time sched pr_t /\
                               t1 <= pr_t <= t1 + K
  t : nat
  GE : t1 + K <= t
  LT : t < t2
  prt : instant
  PR : preemption_time sched prt
  GEprt : t1 <= prt
  LEprt : prt <= t1 + K
  ============================
  t1 + K < t2

subgoal 2 (ID 2625) is:
 prt <= t < t2

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


           by apply leq_ltn_trans with t.

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

1 subgoal (ID 2625)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  K : duration
  H_preemption_time_exists : exists pr_t : instant,
                               preemption_time sched pr_t /\
                               t1 <= pr_t <= t1 + K
  t : nat
  GE : t1 + K <= t
  LT : t < t2
  prt : instant
  PR : preemption_time sched prt
  GEprt : t1 <= prt
  LEprt : prt <= t1 + K
  ============================
  prt <= t < t2

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


      - apply/andP; split; last by done.

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

1 subgoal (ID 2698)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  K : duration
  H_preemption_time_exists : exists pr_t : instant,
                               preemption_time sched pr_t /\
                               t1 <= pr_t <= t1 + K
  t : nat
  GE : t1 + K <= t
  LT : t < t2
  prt : instant
  PR : preemption_time sched prt
  GEprt : t1 <= prt
  LEprt : prt <= t1 + K
  ============================
  prt <= t

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


          by apply leq_trans with (t1 + K).

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

No more subgoals.

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


    Qed.

  End PreemptionTimeAndPriorityInversion.

Preemption Time Exists

In this section we prove that the function [max_length_of_priority_inversion] indeed upper bounds the priority inversion length.
  Section PreemptionTimeExists.

First we prove that if a job with higher-or-equal priority is scheduled at a quiet time [t+1] then this is the first time when this job is scheduled.
    Lemma hp_job_not_scheduled_before_quiet_time:
       jhp t,
        quiet_time arr_seq sched j t.+1
        job_scheduled_at jhp t.+1
        hep_job jhp j
        ~~ job_scheduled_at jhp t.

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

1 subgoal (ID 2494)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  ============================
  forall (jhp : Job) (t : nat),
  quiet_time arr_seq sched j (succn t) ->
  job_scheduled_at jhp (succn t) ->
  hep_job jhp j -> ~~ job_scheduled_at jhp t

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


    Proof.
      intros jhp t QT SCHED1 HP.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2499)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jhp : Job
  t : nat
  QT : quiet_time arr_seq sched j (succn t)
  SCHED1 : job_scheduled_at jhp (succn t)
  HP : hep_job jhp j
  ============================
  ~~ job_scheduled_at jhp t

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



      apply/negP; intros SCHED2.

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

1 subgoal (ID 2521)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jhp : Job
  t : nat
  QT : quiet_time arr_seq sched j (succn t)
  SCHED1 : job_scheduled_at jhp (succn t)
  HP : hep_job jhp j
  SCHED2 : job_scheduled_at jhp t
  ============================
  False

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


      specialize (QT jhp).

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

1 subgoal (ID 2523)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jhp : Job
  t : nat
  QT : arrives_in arr_seq jhp ->
       hep_job jhp j ->
       arrived_before jhp (succn t) -> completed_by sched jhp (succn t)
  SCHED1 : job_scheduled_at jhp (succn t)
  HP : hep_job jhp j
  SCHED2 : job_scheduled_at jhp t
  ============================
  False

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


      feed_n 3 QT; eauto.

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

1 subgoal (ID 2541)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jhp : Job
  t : nat
  QT : completed_by sched jhp (succn t)
  SCHED1 : job_scheduled_at jhp (succn t)
  HP : hep_job jhp j
  SCHED2 : job_scheduled_at jhp t
  ============================
  False

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


      apply completed_implies_not_scheduled in QT; last by done.

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

1 subgoal (ID 2651)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jhp : Job
  t : nat
  QT : ~~ scheduled_at sched jhp (succn t)
  SCHED1 : job_scheduled_at jhp (succn t)
  HP : hep_job jhp j
  SCHED2 : job_scheduled_at jhp t
  ============================
  False

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


        by move: QT ⇒ /negP NSCHED; apply: NSCHED.

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

No more subgoals.

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


    Qed.

Also, we show that lower-priority jobs that are scheduled inside the busy-interval prefix [t1,t2) must arrive before that interval.
    Lemma low_priority_job_arrives_before_busy_interval_prefix:
       jlp t,
        t1 t < t2
        job_scheduled_at jlp t
        ~~ hep_job jlp j
        job_arrival jlp < t1.

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

1 subgoal (ID 2501)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  ============================
  forall (jlp : Job) (t : nat),
  t1 <= t < t2 ->
  job_scheduled_at jlp t -> ~~ hep_job jlp j -> job_arrival jlp < t1

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


    Proof.
      movejlp t /andP [GE LT] SCHED LP.

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

1 subgoal (ID 2545)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  GE : t1 <= t
  LT : t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  ============================
  job_arrival jlp < t1

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


      move: (H_busy_interval_prefix) ⇒ [NEM [QT [NQT HPJ]]].

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

1 subgoal (ID 2577)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  GE : t1 <= t
  LT : t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ============================
  job_arrival jlp < t1

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


      apply negbNE; apply/negP; intros ARR; rewrite -leqNgt in ARR.

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

1 subgoal (ID 2650)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  GE : t1 <= t
  LT : t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ARR : t1 <= job_arrival jlp
  ============================
  False

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


      have SCH:= scheduling_of_any_segment_starts_with_preemption_time _ _ SCHED.

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

1 subgoal (ID 2659)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  GE : t1 <= t
  LT : t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ARR : t1 <= job_arrival jlp
  SCH : exists pt : nat,
          job_arrival jlp <= pt <= t /\
          preemption_time sched pt /\
          (forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t')
  ============================
  False

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


      move: SCH ⇒ [pt [/andP [NEQ1 NEQ2] [PT FA]]].

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

1 subgoal (ID 2731)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  GE : t1 <= t
  LT : t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ARR : t1 <= job_arrival jlp
  pt : nat
  NEQ1 : job_arrival jlp <= pt
  NEQ2 : pt <= t
  PT : preemption_time sched pt
  FA : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  ============================
  False

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


      have NEQ: t1 pt < t2.

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

2 subgoals (ID 2732)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  GE : t1 <= t
  LT : t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ARR : t1 <= job_arrival jlp
  pt : nat
  NEQ1 : job_arrival jlp <= pt
  NEQ2 : pt <= t
  PT : preemption_time sched pt
  FA : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  ============================
  t1 <= pt < t2

subgoal 2 (ID 2734) is:
 False

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


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

1 subgoal (ID 2732)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  GE : t1 <= t
  LT : t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ARR : t1 <= job_arrival jlp
  pt : nat
  NEQ1 : job_arrival jlp <= pt
  NEQ2 : pt <= t
  PT : preemption_time sched pt
  FA : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  ============================
  t1 <= pt < t2

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


apply/andP; split.

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

2 subgoals (ID 2760)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  GE : t1 <= t
  LT : t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ARR : t1 <= job_arrival jlp
  pt : nat
  NEQ1 : job_arrival jlp <= pt
  NEQ2 : pt <= t
  PT : preemption_time sched pt
  FA : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  ============================
  t1 <= pt

subgoal 2 (ID 2761) is:
 pt < t2

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


        apply leq_trans with (job_arrival jlp); by done.

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

1 subgoal (ID 2761)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  GE : t1 <= t
  LT : t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ARR : t1 <= job_arrival jlp
  pt : nat
  NEQ1 : job_arrival jlp <= pt
  NEQ2 : pt <= t
  PT : preemption_time sched pt
  FA : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  ============================
  pt < t2

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


        apply leq_ltn_trans with t; by done.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2734)

subgoal 1 (ID 2734) is:
 False

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


}

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

1 subgoal (ID 2734)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  GE : t1 <= t
  LT : t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ARR : t1 <= job_arrival jlp
  pt : nat
  NEQ1 : job_arrival jlp <= pt
  NEQ2 : pt <= t
  PT : preemption_time sched pt
  FA : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  NEQ : t1 <= pt < t2
  ============================
  False

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


      have LL:= not_quiet_implies_exists_scheduled_hp_job_at_preemption_point pt.

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

1 subgoal (ID 2772)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  GE : t1 <= t
  LT : t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ARR : t1 <= job_arrival jlp
  pt : nat
  NEQ1 : job_arrival jlp <= pt
  NEQ2 : pt <= t
  PT : preemption_time sched pt
  FA : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  NEQ : t1 <= pt < t2
  LL : t1 <= pt < t2 ->
       preemption_time sched pt ->
       exists j_hp : Job,
         arrived_between j_hp t1 t2 /\
         hep_job j_hp j /\ job_scheduled_at j_hp pt
  ============================
  False

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


      feed_n 2 LL; try done.

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

1 subgoal (ID 2784)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  GE : t1 <= t
  LT : t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ARR : t1 <= job_arrival jlp
  pt : nat
  NEQ1 : job_arrival jlp <= pt
  NEQ2 : pt <= t
  PT : preemption_time sched pt
  FA : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  NEQ : t1 <= pt < t2
  LL : exists j_hp : Job,
         arrived_between j_hp t1 t2 /\
         hep_job j_hp j /\ job_scheduled_at j_hp pt
  ============================
  False

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


      move: LL ⇒ [jhp [ARRjhp [HP SCHEDhp]]].

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

1 subgoal (ID 2843)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  GE : t1 <= t
  LT : t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ARR : t1 <= job_arrival jlp
  pt : nat
  NEQ1 : job_arrival jlp <= pt
  NEQ2 : pt <= t
  PT : preemption_time sched pt
  FA : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  NEQ : t1 <= pt < t2
  jhp : Job
  ARRjhp : arrived_between jhp t1 t2
  HP : hep_job jhp j
  SCHEDhp : job_scheduled_at jhp pt
  ============================
  False

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


      feed (FA pt); first (by apply/andP; split).
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2849)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  GE : t1 <= t
  LT : t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ARR : t1 <= job_arrival jlp
  pt : nat
  NEQ1 : job_arrival jlp <= pt
  NEQ2 : pt <= t
  PT : preemption_time sched pt
  FA : job_scheduled_at jlp pt
  NEQ : t1 <= pt < t2
  jhp : Job
  ARRjhp : arrived_between jhp t1 t2
  HP : hep_job jhp j
  SCHEDhp : job_scheduled_at jhp pt
  ============================
  False

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


      move: LP ⇒ /negP LP; apply: LP.

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

1 subgoal (ID 2910)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  GE : t1 <= t
  LT : t < t2
  SCHED : job_scheduled_at jlp t
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ARR : t1 <= job_arrival jlp
  pt : nat
  NEQ1 : job_arrival jlp <= pt
  NEQ2 : pt <= t
  PT : preemption_time sched pt
  FA : job_scheduled_at jlp pt
  NEQ : t1 <= pt < t2
  jhp : Job
  ARRjhp : arrived_between jhp t1 t2
  HP : hep_job jhp j
  SCHEDhp : job_scheduled_at jhp pt
  ============================
  hep_job jlp j

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


        by have ->: jlp = jhp by eapply ideal_proc_model_is_a_uniprocessor_model; eauto.

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

No more subgoals.

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


    Qed.

Moreover, we show that lower-priority jobs that are scheduled inside the busy-interval prefix [t1,t2) must be scheduled before that interval.
    Lemma low_priority_job_scheduled_before_busy_interval_prefix:
       jlp t,
        t1 t < t2
        job_scheduled_at jlp t
        ~~ hep_job jlp j
         t', t' < t1 job_scheduled_at jlp t'.

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

1 subgoal (ID 2508)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  ============================
  forall (jlp : Job) (t : nat),
  t1 <= t < t2 ->
  job_scheduled_at jlp t ->
  ~~ hep_job jlp j -> exists t' : nat, t' < t1 /\ job_scheduled_at jlp t'

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


    Proof.
      movejlp t NEQ SCHED LP; move: (NEQ) ⇒ /andP [GE LT].

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

1 subgoal (ID 2554)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ============================
  exists t' : nat, t' < t1 /\ job_scheduled_at jlp t'

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


      have ARR := low_priority_job_arrives_before_busy_interval_prefix _ _ NEQ SCHED LP.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2563)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  ============================
  exists t' : nat, t' < t1 /\ job_scheduled_at jlp t'

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


       t1.-1; split.

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

2 subgoals (ID 2567)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  ============================
  predn t1 < t1

subgoal 2 (ID 2568) is:
 job_scheduled_at jlp (predn t1)

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


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

1 subgoal (ID 2567)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  ============================
  predn t1 < t1

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


by rewrite prednK; last apply leq_ltn_trans with (job_arrival jlp).
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2568)

subgoal 1 (ID 2568) is:
 job_scheduled_at jlp (predn t1)

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


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

1 subgoal (ID 2568)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  ============================
  job_scheduled_at jlp (predn t1)

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



      move: (H_busy_interval_prefix) ⇒ [NEM [QT [NQT HPJ]]].

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

1 subgoal (ID 2609)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ============================
  job_scheduled_at jlp (predn t1)

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


      have SCHEDST := scheduling_of_any_segment_starts_with_preemption_time _ _ SCHED.

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

1 subgoal (ID 2618)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  SCHEDST : exists pt : nat,
              job_arrival jlp <= pt <= t /\
              preemption_time sched pt /\
              (forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t')
  ============================
  job_scheduled_at jlp (predn t1)

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


      move: SCHEDST ⇒ [pt [NEQpt [PT SCHEDc]]].

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

1 subgoal (ID 2651)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  pt : nat
  NEQpt : job_arrival jlp <= pt <= t
  PT : preemption_time sched pt
  SCHEDc : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  ============================
  job_scheduled_at jlp (predn t1)

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


      have LT2: pt < t1.

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

2 subgoals (ID 2652)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  pt : nat
  NEQpt : job_arrival jlp <= pt <= t
  PT : preemption_time sched pt
  SCHEDc : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  ============================
  pt < t1

subgoal 2 (ID 2654) is:
 job_scheduled_at jlp (predn t1)

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


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

1 subgoal (ID 2652)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  pt : nat
  NEQpt : job_arrival jlp <= pt <= t
  PT : preemption_time sched pt
  SCHEDc : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  ============================
  pt < t1

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


rewrite ltnNge; apply/negP; intros CONTR.

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

1 subgoal (ID 2680)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  pt : nat
  NEQpt : job_arrival jlp <= pt <= t
  PT : preemption_time sched pt
  SCHEDc : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  CONTR : t1 <= pt
  ============================
  False

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


        have NQSCHED := not_quiet_implies_exists_scheduled_hp_job_at_preemption_point pt.

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

1 subgoal (ID 2685)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  pt : nat
  NEQpt : job_arrival jlp <= pt <= t
  PT : preemption_time sched pt
  SCHEDc : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  CONTR : t1 <= pt
  NQSCHED : t1 <= pt < t2 ->
            preemption_time sched pt ->
            exists j_hp : Job,
              arrived_between j_hp t1 t2 /\
              hep_job j_hp j /\ job_scheduled_at j_hp pt
  ============================
  False

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


        feed_n 2 NQSCHED; try done.

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

2 subgoals (ID 2686)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  pt : nat
  NEQpt : job_arrival jlp <= pt <= t
  PT : preemption_time sched pt
  SCHEDc : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  CONTR : t1 <= pt
  NQSCHED : t1 <= pt < t2 ->
            preemption_time sched pt ->
            exists j_hp : Job,
              arrived_between j_hp t1 t2 /\
              hep_job j_hp j /\ job_scheduled_at j_hp pt
  ============================
  t1 <= pt < t2

subgoal 2 (ID 2697) is:
 False

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


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

1 subgoal (ID 2686)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  pt : nat
  NEQpt : job_arrival jlp <= pt <= t
  PT : preemption_time sched pt
  SCHEDc : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  CONTR : t1 <= pt
  NQSCHED : t1 <= pt < t2 ->
            preemption_time sched pt ->
            exists j_hp : Job,
              arrived_between j_hp t1 t2 /\
              hep_job j_hp j /\ job_scheduled_at j_hp pt
  ============================
  t1 <= pt < t2

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


apply/andP; split; first by done.

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

1 subgoal (ID 2772)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  pt : nat
  NEQpt : job_arrival jlp <= pt <= t
  PT : preemption_time sched pt
  SCHEDc : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  CONTR : t1 <= pt
  NQSCHED : t1 <= pt < t2 ->
            preemption_time sched pt ->
            exists j_hp : Job,
              arrived_between j_hp t1 t2 /\
              hep_job j_hp j /\ job_scheduled_at j_hp pt
  ============================
  pt < t2

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


            by apply leq_ltn_trans with t; move: NEQpt ⇒ /andP [_ T].

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

2 subgoals (ID 2697)

subgoal 1 (ID 2697) is:
 False
subgoal 2 (ID 2654) is:
 job_scheduled_at jlp (predn t1)

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


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

1 subgoal (ID 2697)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  pt : nat
  NEQpt : job_arrival jlp <= pt <= t
  PT : preemption_time sched pt
  SCHEDc : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  CONTR : t1 <= pt
  NQSCHED : exists j_hp : Job,
              arrived_between j_hp t1 t2 /\
              hep_job j_hp j /\ job_scheduled_at j_hp pt
  ============================
  False

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


move: NQSCHED ⇒ [jhp [ARRhp [HPhp SCHEDhp]]].

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

1 subgoal (ID 2891)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  pt : nat
  NEQpt : job_arrival jlp <= pt <= t
  PT : preemption_time sched pt
  SCHEDc : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  CONTR : t1 <= pt
  jhp : Job
  ARRhp : arrived_between jhp t1 t2
  HPhp : hep_job jhp j
  SCHEDhp : job_scheduled_at jhp pt
  ============================
  False

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


        specialize (SCHEDc pt).

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

1 subgoal (ID 2893)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  pt : nat
  NEQpt : job_arrival jlp <= pt <= t
  PT : preemption_time sched pt
  SCHEDc : pt <= pt <= t -> job_scheduled_at jlp pt
  CONTR : t1 <= pt
  jhp : Job
  ARRhp : arrived_between jhp t1 t2
  HPhp : hep_job jhp j
  SCHEDhp : job_scheduled_at jhp pt
  ============================
  False

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


        feed SCHEDc; first by apply/andP; split; last move: NEQpt ⇒ /andP [_ T].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2899)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  pt : nat
  NEQpt : job_arrival jlp <= pt <= t
  PT : preemption_time sched pt
  SCHEDc : job_scheduled_at jlp pt
  CONTR : t1 <= pt
  jhp : Job
  ARRhp : arrived_between jhp t1 t2
  HPhp : hep_job jhp j
  SCHEDhp : job_scheduled_at jhp pt
  ============================
  False

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


        move: LP ⇒ /negP LP; apply: LP.

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

1 subgoal (ID 3002)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  pt : nat
  NEQpt : job_arrival jlp <= pt <= t
  PT : preemption_time sched pt
  SCHEDc : job_scheduled_at jlp pt
  CONTR : t1 <= pt
  jhp : Job
  ARRhp : arrived_between jhp t1 t2
  HPhp : hep_job jhp j
  SCHEDhp : job_scheduled_at jhp pt
  ============================
  hep_job jlp j

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


          by have ->: jlp = jhp by eapply ideal_proc_model_is_a_uniprocessor_model; eauto.

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

1 subgoal (ID 2654)

subgoal 1 (ID 2654) is:
 job_scheduled_at jlp (predn t1)

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


      }

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

1 subgoal (ID 2654)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  pt : nat
  NEQpt : job_arrival jlp <= pt <= t
  PT : preemption_time sched pt
  SCHEDc : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  LT2 : pt < t1
  ============================
  job_scheduled_at jlp (predn t1)

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


      apply SCHEDc; apply/andP; split.

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

2 subgoals (ID 3046)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  pt : nat
  NEQpt : job_arrival jlp <= pt <= t
  PT : preemption_time sched pt
  SCHEDc : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  LT2 : pt < t1
  ============================
  pt <= predn t1

subgoal 2 (ID 3047) is:
 predn t1 <= t

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


      - by rewrite -addn1 in LT2; apply subh3 in LT2; rewrite subn1 in LT2.

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

1 subgoal (ID 3047)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  t : nat
  NEQ : t1 <= t < t2
  SCHED : job_scheduled_at jlp t
  LP : ~~ hep_job jlp j
  GE : t1 <= t
  LT : t < t2
  ARR : job_arrival jlp < t1
  NEM : t1 < t2
  QT : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  pt : nat
  NEQpt : job_arrival jlp <= pt <= t
  PT : preemption_time sched pt
  SCHEDc : forall t' : nat, pt <= t' <= t -> job_scheduled_at jlp t'
  LT2 : pt < t1
  ============================
  predn t1 <= t

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


      - by apply leq_trans with t1; first apply leq_pred.
(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

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


    Qed.

Thus, there must be a preemption time in the interval [t1, t1 + max_priority_inversion t1]. That is, if a job with higher-or-equal priority is scheduled at time instant t1, then t1 is a preemption time. Otherwise, if a job with lower priority is scheduled at time t1, then this jobs also should be scheduled before the beginning of the busy interval. So, the next preemption time will be no more than [max_priority_inversion t1] time units later.
We proceed by doing a case analysis.
    Section CaseAnalysis.

(1) Case when the schedule is idle at time [t1].
      Section Case1.

Assume that the schedule is idle at time [t1].
        Hypothesis H_is_idle : is_idle sched t1.

Then time instant [t1] is a preemption time.
        Lemma preemption_time_exists_case1:
           pr_t,
            preemption_time sched pr_t
            t1 pr_t t1 + max_length_of_priority_inversion j t1.

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

1 subgoal (ID 2514)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_is_idle : is_idle sched t1
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


        Proof.
          set (service := service sched).

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

1 subgoal (ID 2519)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_is_idle : is_idle sched t1
  service := service.service sched : Job -> instant -> nat
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


          move: (H_valid_model_with_bounded_nonpreemptive_segments) ⇒ CORR.

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

1 subgoal (ID 2521)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_is_idle : is_idle sched t1
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


          move: (H_busy_interval_prefix) ⇒ [NEM [QT1 [NQT HPJ]]].

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

1 subgoal (ID 2553)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_is_idle : is_idle sched t1
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  NEM : t1 < t2
  QT1 : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


           t1; split.

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

2 subgoals (ID 2557)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_is_idle : is_idle sched t1
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  NEM : t1 < t2
  QT1 : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ============================
  preemption_time sched t1

subgoal 2 (ID 2558) is:
 t1 <= t1 <= t1 + max_length_of_priority_inversion j t1

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


          - by rewrite /preemption_time (eqbool_to_eqprop H_is_idle).

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

1 subgoal (ID 2558)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  H_is_idle : is_idle sched t1
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  NEM : t1 < t2
  QT1 : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ============================
  t1 <= t1 <= t1 + max_length_of_priority_inversion j t1

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


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

No more subgoals.

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


        Qed.

      End Case1.

(2) Case when a job with higher-or-equal priority is scheduled at time [t1].
      Section Case2.

Assume that a job [jhp] with higher-or-equal priority is scheduled at time [t1].
        Variable jhp : Job.
        Hypothesis H_jhp_is_scheduled : scheduled_at sched jhp t1.
        Hypothesis H_jhp_hep_priority : hep_job jhp j.

Then time instant [t1] is a preemption time.
        Lemma preemption_time_exists_case2:
           pr_t,
            preemption_time sched pr_t
            t1 pr_t t1 + max_length_of_priority_inversion j t1.

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

1 subgoal (ID 2518)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jhp : Job
  H_jhp_is_scheduled : scheduled_at sched jhp t1
  H_jhp_hep_priority : hep_job jhp j
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


        Proof.
          set (service := service sched).

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

1 subgoal (ID 2523)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jhp : Job
  H_jhp_is_scheduled : scheduled_at sched jhp t1
  H_jhp_hep_priority : hep_job jhp j
  service := service.service sched : Job -> instant -> nat
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


          move: (H_valid_model_with_bounded_nonpreemptive_segments) ⇒ CORR.

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

1 subgoal (ID 2525)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jhp : Job
  H_jhp_is_scheduled : scheduled_at sched jhp t1
  H_jhp_hep_priority : hep_job jhp j
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


          move: (H_busy_interval_prefix) ⇒ [NEM [QT1 [NQT HPJ]]].

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

1 subgoal (ID 2557)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jhp : Job
  H_jhp_is_scheduled : scheduled_at sched jhp t1
  H_jhp_hep_priority : hep_job jhp j
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  NEM : t1 < t2
  QT1 : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


           t1; split; last first.

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

2 subgoals (ID 2562)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jhp : Job
  H_jhp_is_scheduled : scheduled_at sched jhp t1
  H_jhp_hep_priority : hep_job jhp j
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  NEM : t1 < t2
  QT1 : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ============================
  t1 <= t1 <= t1 + max_length_of_priority_inversion j t1

subgoal 2 (ID 2561) is:
 preemption_time sched t1

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


          apply/andP; split; [by done | by rewrite leq_addr].

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

1 subgoal (ID 2561)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jhp : Job
  H_jhp_is_scheduled : scheduled_at sched jhp t1
  H_jhp_hep_priority : hep_job jhp j
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  NEM : t1 < t2
  QT1 : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ============================
  preemption_time sched t1

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


          destruct t1.

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

2 subgoals (ID 2630)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j 0 t2
  jhp : Job
  H_jhp_is_scheduled : scheduled_at sched jhp 0
  H_jhp_hep_priority : hep_job jhp j
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  NEM : 0 < t2
  QT1 : quiet_time arr_seq sched j 0
  NQT : forall t : nat, 0 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : 0 <= job_arrival j < t2
  ============================
  preemption_time sched 0

subgoal 2 (ID 2640) is:
 preemption_time sched (succn i)

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


          - eapply zero_is_pt; eauto 2.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2640)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  i : nat
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j (succn i) t2
  jhp : Job
  H_jhp_is_scheduled : scheduled_at sched jhp (succn i)
  H_jhp_hep_priority : hep_job jhp j
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  NEM : succn i < t2
  QT1 : quiet_time arr_seq sched j (succn i)
  NQT : forall t : nat, succn i < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : i < job_arrival j < t2
  ============================
  preemption_time sched (succn i)

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


          - eapply hp_job_not_scheduled_before_quiet_time in QT1; eauto 2.

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

1 subgoal (ID 2652)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  i : nat
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j (succn i) t2
  jhp : Job
  H_jhp_is_scheduled : scheduled_at sched jhp (succn i)
  H_jhp_hep_priority : hep_job jhp j
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  NEM : succn i < t2
  NQT : forall t : nat, succn i < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : i < job_arrival j < t2
  QT1 : ~~ job_scheduled_at jhp i
  ============================
  preemption_time sched (succn i)

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


              by eapply first_moment_is_pt with (j0 := jhp); eauto 2.

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

No more subgoals.

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


        Qed.

      End Case2.

(3) Case when a job with lower priority is scheduled at time [t1].
      Section Case3.

Assume that a job [jhp] with lower priority is scheduled at time [t1].
        Variable jlp : Job.
        Hypothesis H_jlp_is_scheduled : scheduled_at sched jlp t1.
        Hypothesis H_jlp_low_priority : ~~ hep_job jlp j.

To prove the lemma in this case we need a few auxiliary facts about the first preemption point of job [jlp].
        Section FirstPreemptionPointOfjlp.

Let's denote the progress of job [jlp] at time [t1] as [progr_t1].
          Let progr_t1 := service sched jlp t1.

Consider the first preemption point of job [jlp] after [progr_t1].
          Variable fpt : instant.
          Hypothesis H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt).
          Hypothesis H_fpt_is_first_preemption_point:
             ρ,
              progr_t1 ρ progr_t1 + (job_max_nonpreemptive_segment jlp - ε)
              job_preemptable jlp ρ
              service sched jlp t1 + fpt ρ.

For convenience we also assume the following inequality holds.
          Hypothesis H_progr_le_max_nonp_segment:
            progr_t1 progr_t1 + fpt progr_t1 + (job_max_nonpreemptive_segment jlp - ε).

First we show that [fpt] is indeed the first preemption point after [progr_t1].
          Lemma no_intermediate_preemption_point:
             ρ,
              progr_t1 ρ < progr_t1 + fpt
              ~~ job_preemptable jlp ρ.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2534)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ============================
  forall ρ : nat, progr_t1 <= ρ < progr_t1 + fpt -> ~~ job_preemptable jlp ρ

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


          Proof.
            moveprog /andP [GE LT].

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

1 subgoal (ID 2575)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  prog : nat
  GE : progr_t1 <= prog
  LT : prog < progr_t1 + fpt
  ============================
  ~~ job_preemptable jlp prog

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


            apply/negP; intros PPJ.

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

1 subgoal (ID 2597)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  prog : nat
  GE : progr_t1 <= prog
  LT : prog < progr_t1 + fpt
  PPJ : job_preemptable jlp prog
  ============================
  False

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


            move: H_fpt_is_first_preemption_pointK; specialize (K prog).
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2601)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  prog : nat
  GE : progr_t1 <= prog
  LT : prog < progr_t1 + fpt
  PPJ : job_preemptable jlp prog
  K : progr_t1 <= prog <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε) ->
      job_preemptable jlp prog -> service sched jlp t1 + fpt <= prog
  ============================
  False

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


            feed_n 2 K; first (apply/andP; split); try done.

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

2 subgoals (ID 2640)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  prog : nat
  GE : progr_t1 <= prog
  LT : prog < progr_t1 + fpt
  PPJ : job_preemptable jlp prog
  K : progr_t1 <= prog <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε) ->
      job_preemptable jlp prog -> service sched jlp t1 + fpt <= prog
  ============================
  prog <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε)

subgoal 2 (ID 2613) is:
 False

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


            apply leq_trans with (service sched jlp t1 + fpt).

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

3 subgoals (ID 2692)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  prog : nat
  GE : progr_t1 <= prog
  LT : prog < progr_t1 + fpt
  PPJ : job_preemptable jlp prog
  K : progr_t1 <= prog <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε) ->
      job_preemptable jlp prog -> service sched jlp t1 + fpt <= prog
  ============================
  prog <= service sched jlp t1 + fpt

subgoal 2 (ID 2693) is:
 service sched jlp t1 + fpt <=
 progr_t1 + (job_max_nonpreemptive_segment jlp - ε)
subgoal 3 (ID 2613) is:
 False

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


            + by apply ltnW.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2693)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  prog : nat
  GE : progr_t1 <= prog
  LT : prog < progr_t1 + fpt
  PPJ : job_preemptable jlp prog
  K : progr_t1 <= prog <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε) ->
      job_preemptable jlp prog -> service sched jlp t1 + fpt <= prog
  ============================
  service sched jlp t1 + fpt <=
  progr_t1 + (job_max_nonpreemptive_segment jlp - ε)

subgoal 2 (ID 2613) is:
 False

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


            + by move: H_progr_le_max_nonp_segment ⇒ /andP [_ LL].

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

1 subgoal (ID 2613)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  prog : nat
  GE : progr_t1 <= prog
  LT : prog < progr_t1 + fpt
  PPJ : job_preemptable jlp prog
  K : service sched jlp t1 + fpt <= prog
  ============================
  False

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


                by move: K; rewrite leqNgt; move ⇒ /negP NLT; apply: NLT.

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

No more subgoals.

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


          Qed.

Thanks to the fact that the scheduler respects the notion of preemption points we show that [jlp] is continuously scheduled in time interval [t1, t1 + fpt).
          Lemma continuously_scheduled_between_preemption_points:
             t',
              t1 t' < t1 + fpt
              job_scheduled_at jlp t'.

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

1 subgoal (ID 2536)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ============================
  forall t' : nat, t1 <= t' < t1 + fpt -> job_scheduled_at jlp t'

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


          Proof.
             move: (H_valid_model_with_bounded_nonpreemptive_segments) ⇒ CORR.

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

1 subgoal (ID 2538)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  ============================
  forall t' : nat, t1 <= t' < t1 + fpt -> job_scheduled_at jlp t'

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


             move: (H_jlp_is_scheduled) ⇒ ARRs; apply H_jobs_come_from_arrival_sequence in ARRs.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2541)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  ARRs : arrives_in arr_seq jlp
  ============================
  forall t' : nat, t1 <= t' < t1 + fpt -> job_scheduled_at jlp t'

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



            movet' /andP [GE LT].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2583)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  ARRs : arrives_in arr_seq jlp
  t' : nat
  GE : t1 <= t'
  LT : t' < t1 + fpt
  ============================
  job_scheduled_at jlp t'

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


              have Fact: Δ, t' = t1 + Δ.

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

2 subgoals (ID 2588)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  ARRs : arrives_in arr_seq jlp
  t' : nat
  GE : t1 <= t'
  LT : t' < t1 + fpt
  ============================
  exists Δ : nat, t' = t1 + Δ

subgoal 2 (ID 2590) is:
 job_scheduled_at jlp t'

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


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

1 subgoal (ID 2588)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  ARRs : arrives_in arr_seq jlp
  t' : nat
  GE : t1 <= t'
  LT : t' < t1 + fpt
  ============================
  exists Δ : nat, t' = t1 + Δ

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


by (t' - t1); apply/eqP; rewrite eq_sym; apply/eqP; rewrite subnKC.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2590)

subgoal 1 (ID 2590) is:
 job_scheduled_at jlp t'

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


}

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

1 subgoal (ID 2590)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  ARRs : arrives_in arr_seq jlp
  t' : nat
  GE : t1 <= t'
  LT : t' < t1 + fpt
  Fact : exists Δ : nat, t' = t1 + Δ
  ============================
  job_scheduled_at jlp t'

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


              move: Fact ⇒ [Δ EQ]; subst t'.

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

1 subgoal (ID 2707)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  ARRs : arrives_in arr_seq jlp
  Δ : nat
  LT : t1 + Δ < t1 + fpt
  GE : t1 <= t1 + Δ
  ============================
  job_scheduled_at jlp (t1 + Δ)

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


              have NPPJ := @no_intermediate_preemption_point (@service _ _ _ sched jlp (t1 + Δ)).

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

1 subgoal (ID 2718)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  ARRs : arrives_in arr_seq jlp
  Δ : nat
  LT : t1 + Δ < t1 + fpt
  GE : t1 <= t1 + Δ
  NPPJ : forall p : ProcessorState Job (processor_state Job),
         progr_t1 <= service sched jlp (t1 + Δ) < progr_t1 + fpt ->
         ~~ job_preemptable jlp (service sched jlp (t1 + Δ))
  ============================
  job_scheduled_at jlp (t1 + Δ)

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


              apply proj1 in CORR; specialize (CORR jlp ARRs).

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

1 subgoal (ID 2722)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : job_cannot_become_nonpreemptive_before_execution jlp /\
         job_cannot_be_nonpreemptive_after_completion jlp /\
         not_preemptive_implies_scheduled sched jlp /\
         execution_starts_with_preemption_point sched jlp
  ARRs : arrives_in arr_seq jlp
  Δ : nat
  LT : t1 + Δ < t1 + fpt
  GE : t1 <= t1 + Δ
  NPPJ : forall p : ProcessorState Job (processor_state Job),
         progr_t1 <= service sched jlp (t1 + Δ) < progr_t1 + fpt ->
         ~~ job_preemptable jlp (service sched jlp (t1 + Δ))
  ============================
  job_scheduled_at jlp (t1 + Δ)

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


              move: CORR ⇒ [_ [_ [T _] ]].

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

1 subgoal (ID 2756)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  Δ : nat
  LT : t1 + Δ < t1 + fpt
  GE : t1 <= t1 + Δ
  NPPJ : forall p : ProcessorState Job (processor_state Job),
         progr_t1 <= service sched jlp (t1 + Δ) < progr_t1 + fpt ->
         ~~ job_preemptable jlp (service sched jlp (t1 + Δ))
  T : not_preemptive_implies_scheduled sched jlp
  ============================
  job_scheduled_at jlp (t1 + Δ)

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


              apply T; apply: NPPJ; apply/andP; split.

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

2 subgoals (ID 2792)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  Δ : nat
  LT : t1 + Δ < t1 + fpt
  GE : t1 <= t1 + Δ
  T : not_preemptive_implies_scheduled sched jlp
  ============================
  progr_t1 <= service sched jlp (t1 + Δ)

subgoal 2 (ID 2793) is:
 service sched jlp (t1 + Δ) < progr_t1 + fpt

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


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

1 subgoal (ID 2792)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  Δ : nat
  LT : t1 + Δ < t1 + fpt
  GE : t1 <= t1 + Δ
  T : not_preemptive_implies_scheduled sched jlp
  ============================
  progr_t1 <= service sched jlp (t1 + Δ)

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


by apply service_monotonic; rewrite leq_addr.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2793)

subgoal 1 (ID 2793) is:
 service sched jlp (t1 + Δ) < progr_t1 + fpt

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


}

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

1 subgoal (ID 2793)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  Δ : nat
  LT : t1 + Δ < t1 + fpt
  GE : t1 <= t1 + Δ
  T : not_preemptive_implies_scheduled sched jlp
  ============================
  service sched jlp (t1 + Δ) < progr_t1 + fpt

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


              rewrite /service -(service_during_cat _ _ _ t1).

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

2 subgoals (ID 2830)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  Δ : nat
  LT : t1 + Δ < t1 + fpt
  GE : t1 <= t1 + Δ
  T : not_preemptive_implies_scheduled sched jlp
  ============================
  service_during sched jlp 0 t1 + service_during sched jlp t1 (t1 + Δ) <
  progr_t1 + fpt

subgoal 2 (ID 2831) is:
 0 <= t1 <= t1 + Δ

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


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

1 subgoal (ID 2830)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  Δ : nat
  LT : t1 + Δ < t1 + fpt
  GE : t1 <= t1 + Δ
  T : not_preemptive_implies_scheduled sched jlp
  ============================
  service_during sched jlp 0 t1 + service_during sched jlp t1 (t1 + Δ) <
  progr_t1 + fpt

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


rewrite ltn_add2l; rewrite ltn_add2l in LT.

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

1 subgoal (ID 2889)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  Δ : nat
  GE : t1 <= t1 + Δ
  T : not_preemptive_implies_scheduled sched jlp
  LT : Δ < fpt
  ============================
  service_during sched jlp t1 (t1 + Δ) < fpt

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


                apply leq_ltn_trans with Δ; last by done.

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

1 subgoal (ID 2890)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  Δ : nat
  GE : t1 <= t1 + Δ
  T : not_preemptive_implies_scheduled sched jlp
  LT : Δ < fpt
  ============================
  service_during sched jlp t1 (t1 + Δ) <= Δ

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


                rewrite -{2}(sum_of_ones t1 Δ).

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

1 subgoal (ID 2893)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  Δ : nat
  GE : t1 <= t1 + Δ
  T : not_preemptive_implies_scheduled sched jlp
  LT : Δ < fpt
  ============================
  service_during sched jlp t1 (t1 + Δ) <= \sum_(t1 <= x < t1 + Δ) 1

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


                rewrite leq_sum //; intros t _.

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

1 subgoal (ID 2931)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  Δ : nat
  GE : t1 <= t1 + Δ
  T : not_preemptive_implies_scheduled sched jlp
  LT : Δ < fpt
  t : nat
  ============================
  service_at sched jlp t <= 1

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


                apply service_at_most_one.

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

1 subgoal (ID 2935)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  Δ : nat
  GE : t1 <= t1 + Δ
  T : not_preemptive_implies_scheduled sched jlp
  LT : Δ < fpt
  t : nat
  ============================
  unit_service_proc_model (processor_state Job)

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


                  by apply ideal_proc_model_provides_unit_service.

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

1 subgoal (ID 2831)

subgoal 1 (ID 2831) is:
 0 <= t1 <= t1 + Δ

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


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

1 subgoal (ID 2831)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  Δ : nat
  LT : t1 + Δ < t1 + fpt
  GE : t1 <= t1 + Δ
  T : not_preemptive_implies_scheduled sched jlp
  ============================
  0 <= t1 <= t1 + Δ

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



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

1 subgoal (ID 2831)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  Δ : nat
  LT : t1 + Δ < t1 + fpt
  GE : t1 <= t1 + Δ
  T : not_preemptive_implies_scheduled sched jlp
  ============================
  0 <= t1 <= t1 + Δ

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


by apply/andP; split; [done | rewrite leq_addr].
(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

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


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

No more subgoals.

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



          Qed.

Thus, job [jlp] reaches its preemption point at time instant [t1 + fpt], which implies that time instant [t1 + fpt] is a preemption time.
          Lemma first_preemption_time:
            preemption_time sched (t1 + fpt).

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

1 subgoal (ID 2539)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ============================
  preemption_time sched (t1 + fpt)

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


          Proof.
            rewrite /preemption_time.

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

1 subgoal (ID 2544)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ============================
  match sched (t1 + fpt) with
  | Some j0 => job_preemptable j0 (service sched j0 (t1 + fpt))
  | None => true
  end

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


            move: (H_valid_model_with_bounded_nonpreemptive_segments) ⇒ CORR.

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

1 subgoal (ID 2546)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  ============================
  match sched (t1 + fpt) with
  | Some j0 => job_preemptable j0 (service sched j0 (t1 + fpt))
  | None => true
  end

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


            ideal_proc_model_sched_case_analysis_eq sched (t1 + fpt) s'; try done.

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

1 subgoal (ID 2642)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : #|[pred x |
                   let
                   'FiniteQuant.Quantified F :=
                    FiniteQuant.ex (T:=Core)
                      (, scheduled_on s' (sched (t1 + fpt)) x) x x in F]| <>
               0
  ============================
  match sched (t1 + fpt) with
  | Some j0 => job_preemptable j0 (service sched j0 (t1 + fpt))
  | None => true
  end

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


            clear EqSched_s'; move: (Sched_s'); rewrite scheduled_at_def;
              move ⇒ /eqP EqSched_s'; rewrite EqSched_s'.

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

1 subgoal (ID 2711)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  ============================
  job_preemptable s' (service sched s' (t1 + fpt))

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


            destruct (jlp == s') eqn: EQ.

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

2 subgoals (ID 2722)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = true
  ============================
  job_preemptable s' (service sched s' (t1 + fpt))

subgoal 2 (ID 2723) is:
 job_preemptable s' (service sched s' (t1 + fpt))

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


            - move: EQ ⇒ /eqP EQ; subst s'.

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

2 subgoals (ID 2767)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  EqSched_s' : sched (t1 + fpt) = Some jlp
  Sched_s' : scheduled_at sched jlp (t1 + fpt)
  ============================
  job_preemptable jlp (service sched jlp (t1 + fpt))

subgoal 2 (ID 2723) is:
 job_preemptable s' (service sched s' (t1 + fpt))

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


              rewrite /service -(service_during_cat _ _ _ t1); last first.

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

3 subgoals (ID 2797)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  EqSched_s' : sched (t1 + fpt) = Some jlp
  Sched_s' : scheduled_at sched jlp (t1 + fpt)
  ============================
  0 <= t1 <= t1 + fpt

subgoal 2 (ID 2796) is:
 job_preemptable jlp
   (service_during sched jlp 0 t1 + service_during sched jlp t1 (t1 + fpt))
subgoal 3 (ID 2723) is:
 job_preemptable s' (service sched s' (t1 + fpt))

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


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

1 subgoal (ID 2797)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  EqSched_s' : sched (t1 + fpt) = Some jlp
  Sched_s' : scheduled_at sched jlp (t1 + fpt)
  ============================
  0 <= t1 <= t1 + fpt

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


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

2 subgoals (ID 2796)

subgoal 1 (ID 2796) is:
 job_preemptable jlp
   (service_during sched jlp 0 t1 + service_during sched jlp t1 (t1 + fpt))
subgoal 2 (ID 2723) is:
 job_preemptable s' (service sched s' (t1 + fpt))

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


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

2 subgoals (ID 2796)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  EqSched_s' : sched (t1 + fpt) = Some jlp
  Sched_s' : scheduled_at sched jlp (t1 + fpt)
  ============================
  job_preemptable jlp
    (service_during sched jlp 0 t1 + service_during sched jlp t1 (t1 + fpt))

subgoal 2 (ID 2723) is:
 job_preemptable s' (service sched s' (t1 + fpt))

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


             
              have ->: service_during sched jlp t1 (t1 + fpt) = fpt.

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

3 subgoals (ID 2834)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  EqSched_s' : sched (t1 + fpt) = Some jlp
  Sched_s' : scheduled_at sched jlp (t1 + fpt)
  ============================
  service_during sched jlp t1 (t1 + fpt) = fpt

subgoal 2 (ID 2839) is:
 job_preemptable jlp (service_during sched jlp 0 t1 + fpt)
subgoal 3 (ID 2723) is:
 job_preemptable s' (service sched s' (t1 + fpt))

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


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

1 subgoal (ID 2834)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  EqSched_s' : sched (t1 + fpt) = Some jlp
  Sched_s' : scheduled_at sched jlp (t1 + fpt)
  ============================
  service_during sched jlp t1 (t1 + fpt) = fpt

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


rewrite -{2}(sum_of_ones t1 fpt) /service_during.

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

1 subgoal (ID 2848)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  EqSched_s' : sched (t1 + fpt) = Some jlp
  Sched_s' : scheduled_at sched jlp (t1 + fpt)
  ============================
  \sum_(t1 <= t < t1 + fpt) service_at sched jlp t =
  \sum_(t1 <= x < t1 + fpt) 1

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


                apply/eqP; rewrite eqn_leq //; apply/andP; split.

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

2 subgoals (ID 2955)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  EqSched_s' : sched (t1 + fpt) = Some jlp
  Sched_s' : scheduled_at sched jlp (t1 + fpt)
  ============================
  \sum_(t1 <= t < t1 + fpt) service_at sched jlp t <=
  \sum_(t1 <= x < t1 + fpt) 1

subgoal 2 (ID 2956) is:
 \sum_(t1 <= x < t1 + fpt) 1 <=
 \sum_(t1 <= t < t1 + fpt) service_at sched jlp t

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


                + rewrite leq_sum //; intros t _.

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

2 subgoals (ID 2994)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  EqSched_s' : sched (t1 + fpt) = Some jlp
  Sched_s' : scheduled_at sched jlp (t1 + fpt)
  t : nat
  ============================
  service_at sched jlp t <= 1

subgoal 2 (ID 2956) is:
 \sum_(t1 <= x < t1 + fpt) 1 <=
 \sum_(t1 <= t < t1 + fpt) service_at sched jlp t

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


                  apply service_at_most_one.

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

2 subgoals (ID 2998)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  EqSched_s' : sched (t1 + fpt) = Some jlp
  Sched_s' : scheduled_at sched jlp (t1 + fpt)
  t : nat
  ============================
  unit_service_proc_model (processor_state Job)

subgoal 2 (ID 2956) is:
 \sum_(t1 <= x < t1 + fpt) 1 <=
 \sum_(t1 <= t < t1 + fpt) service_at sched jlp t

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


                    by apply ideal_proc_model_provides_unit_service.

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

1 subgoal (ID 2956)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  EqSched_s' : sched (t1 + fpt) = Some jlp
  Sched_s' : scheduled_at sched jlp (t1 + fpt)
  ============================
  \sum_(t1 <= x < t1 + fpt) 1 <=
  \sum_(t1 <= t < t1 + fpt) service_at sched jlp t

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


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

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

1 subgoal (ID 3022)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  EqSched_s' : sched (t1 + fpt) = Some jlp
  Sched_s' : scheduled_at sched jlp (t1 + fpt)
  ============================
  \sum_(t1 <= i < t1 + fpt | (t1 <= i < t1 + fpt) && true) 1 <=
  \sum_(t1 <= i < t1 + fpt | (t1 <= i < t1 + fpt) && true)
     service_at sched jlp i

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


                  rewrite leq_sum //.

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

1 subgoal (ID 3031)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  EqSched_s' : sched (t1 + fpt) = Some jlp
  Sched_s' : scheduled_at sched jlp (t1 + fpt)
  ============================
  forall i : nat, (t1 <= i < t1 + fpt) && true -> 0 < service_at sched jlp i

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


                  movex /andP [HYP _].

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

1 subgoal (ID 3098)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  EqSched_s' : sched (t1 + fpt) = Some jlp
  Sched_s' : scheduled_at sched jlp (t1 + fpt)
  x : nat
  HYP : t1 <= x < t1 + fpt
  ============================
  0 < service_at sched jlp x

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


                    by rewrite lt0b -scheduled_at_def; apply continuously_scheduled_between_preemption_points.

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

2 subgoals (ID 2839)

subgoal 1 (ID 2839) is:
 job_preemptable jlp (service_during sched jlp 0 t1 + fpt)
subgoal 2 (ID 2723) is:
 job_preemptable s' (service sched s' (t1 + fpt))

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


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

2 subgoals (ID 2839)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  EqSched_s' : sched (t1 + fpt) = Some jlp
  Sched_s' : scheduled_at sched jlp (t1 + fpt)
  ============================
  job_preemptable jlp (service_during sched jlp 0 t1 + fpt)

subgoal 2 (ID 2723) is:
 job_preemptable s' (service sched s' (t1 + fpt))

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


by done.

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

1 subgoal (ID 2723)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = false
  ============================
  job_preemptable s' (service sched s' (t1 + fpt))

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


            - case: (posnP fpt) ⇒ [ZERO|POS].

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

2 subgoals (ID 3120)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = false
  ZERO : fpt = 0
  ============================
  job_preemptable s' (service sched s' (t1 + fpt))

subgoal 2 (ID 3121) is:
 job_preemptable s' (service sched s' (t1 + fpt))

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


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

1 subgoal (ID 3120)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = false
  ZERO : fpt = 0
  ============================
  job_preemptable s' (service sched s' (t1 + fpt))

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


subst fpt.

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

1 subgoal (ID 3135)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + 0 <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + 0 <= ρ
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + 0)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  EqSched_s' : sched (t1 + 0) = Some s'
  Sched_s' : scheduled_at sched s' (t1 + 0)
  EQ : (jlp == s') = false
  ============================
  job_preemptable s' (service sched s' (t1 + 0))

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


                exfalso; move: EQ ⇒ /negP EQ; apply: EQ.

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

1 subgoal (ID 3184)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + 0 <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + 0 <= ρ
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + 0)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  EqSched_s' : sched (t1 + 0) = Some s'
  Sched_s' : scheduled_at sched s' (t1 + 0)
  ============================
  jlp == s'

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


                move: H_jlp_is_scheduled; rewrite scheduled_at_def; move ⇒ /eqP SCHED2.

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

1 subgoal (ID 3226)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + 0 <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + 0 <= ρ
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + 0)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  EqSched_s' : sched (t1 + 0) = Some s'
  Sched_s' : scheduled_at sched s' (t1 + 0)
  SCHED2 : sched t1 = Some jlp
  ============================
  jlp == s'

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


                rewrite addn0 in EqSched_s'; rewrite EqSched_s' in SCHED2.

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

1 subgoal (ID 3327)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + 0 <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + 0 <= ρ
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + 0)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + 0)
  EqSched_s' : sched t1 = Some s'
  SCHED2 : Some s' = Some jlp
  ============================
  jlp == s'

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


                  by inversion SCHED2.

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

1 subgoal (ID 3121)

subgoal 1 (ID 3121) is:
 job_preemptable s' (service sched s' (t1 + fpt))

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


              }

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

1 subgoal (ID 3121)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = false
  POS : 0 < fpt
  ============================
  job_preemptable s' (service sched s' (t1 + fpt))

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


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

1 subgoal (ID 3121)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = false
  POS : 0 < fpt
  ============================
  job_preemptable s' (service sched s' (t1 + fpt))

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


have EX: sm, sm.+1 = fpt.

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

2 subgoals (ID 3349)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = false
  POS : 0 < fpt
  ============================
  exists sm : nat, succn sm = fpt

subgoal 2 (ID 3351) is:
 job_preemptable s' (service sched s' (t1 + fpt))

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


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

1 subgoal (ID 3349)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = false
  POS : 0 < fpt
  ============================
  exists sm : nat, succn sm = fpt

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


fpt.-1.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 3353)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = false
  POS : 0 < fpt
  ============================
  succn (predn fpt) = fpt

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


ssrlia.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 3351)

subgoal 1 (ID 3351) is:
 job_preemptable s' (service sched s' (t1 + fpt))

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


}

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

1 subgoal (ID 3351)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = false
  POS : 0 < fpt
  EX : exists sm : nat, succn sm = fpt
  ============================
  job_preemptable s' (service sched s' (t1 + fpt))

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


                destruct EX as [sm EQ2].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 5218)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = false
  POS : 0 < fpt
  sm : nat
  EQ2 : succn sm = fpt
  ============================
  job_preemptable s' (service sched s' (t1 + fpt))

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


rewrite -EQ2.

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

1 subgoal (ID 5220)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = false
  POS : 0 < fpt
  sm : nat
  EQ2 : succn sm = fpt
  ============================
  job_preemptable s' (service sched s' (t1 + succn sm))

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


                rewrite addnS.

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

1 subgoal (ID 5224)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = false
  POS : 0 < fpt
  sm : nat
  EQ2 : succn sm = fpt
  ============================
  job_preemptable s' (service sched s' (succn (t1 + sm)))

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


                move: ((proj1 CORR) s' (H_jobs_come_from_arrival_sequence _ _ Sched_s')) ⇒ T.

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

1 subgoal (ID 5232)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = false
  POS : 0 < fpt
  sm : nat
  EQ2 : succn sm = fpt
  T : job_cannot_become_nonpreemptive_before_execution s' /\
      job_cannot_be_nonpreemptive_after_completion s' /\
      not_preemptive_implies_scheduled sched s' /\
      execution_starts_with_preemption_point sched s'
  ============================
  job_preemptable s' (service sched s' (succn (t1 + sm)))

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


                apply T; clear T.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 5247)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = false
  POS : 0 < fpt
  sm : nat
  EQ2 : succn sm = fpt
  ============================
  ~~ scheduled_at sched s' (t1 + sm)

subgoal 2 (ID 5248) is:
 scheduled_at sched s' (succn (t1 + sm))

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


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

2 subgoals (ID 5270)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = false
  POS : 0 < fpt
  sm : nat
  EQ2 : succn sm = fpt
  CONTR : scheduled_at sched s' (t1 + sm)
  ============================
  False

subgoal 2 (ID 5248) is:
 scheduled_at sched s' (succn (t1 + sm))

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


                move: EQ ⇒ /negP EQ; apply: EQ.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 5318)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  POS : 0 < fpt
  sm : nat
  EQ2 : succn sm = fpt
  CONTR : scheduled_at sched s' (t1 + sm)
  ============================
  jlp == s'

subgoal 2 (ID 5248) is:
 scheduled_at sched s' (succn (t1 + sm))

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



                move: (continuously_scheduled_between_preemption_points (t1 + sm)) ⇒ SCHEDs0.

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

2 subgoals (ID 5320)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  POS : 0 < fpt
  sm : nat
  EQ2 : succn sm = fpt
  CONTR : scheduled_at sched s' (t1 + sm)
  SCHEDs0 : t1 <= t1 + sm < t1 + fpt -> job_scheduled_at jlp (t1 + sm)
  ============================
  jlp == s'

subgoal 2 (ID 5248) is:
 scheduled_at sched s' (succn (t1 + sm))

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


                feed SCHEDs0; first by apply/andP; split; [rewrite leq_addr | rewrite -EQ2 addnS].

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

2 subgoals (ID 5326)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  POS : 0 < fpt
  sm : nat
  EQ2 : succn sm = fpt
  CONTR : scheduled_at sched s' (t1 + sm)
  SCHEDs0 : job_scheduled_at jlp (t1 + sm)
  ============================
  jlp == s'

subgoal 2 (ID 5248) is:
 scheduled_at sched s' (succn (t1 + sm))

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


                apply/eqP; eapply ideal_proc_model_is_a_uniprocessor_model; eauto 2.

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

1 subgoal (ID 5248)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  s' : Job
  Sched_s' : scheduled_at sched s' (t1 + fpt)
  EqSched_s' : sched (t1 + fpt) = Some s'
  EQ : (jlp == s') = false
  POS : 0 < fpt
  sm : nat
  EQ2 : succn sm = fpt
  ============================
  scheduled_at sched s' (succn (t1 + sm))

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


                  by rewrite -addnS EQ2.
(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

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


              }

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

No more subgoals.

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


          Qed.

And since [fpt <= max_length_of_priority_inversion j t1], [t1 <= t1 + fpt <= t1 + max_length_of_priority_inversion j t1].
          Lemma preemption_time_le_max_len_of_priority_inversion:
            t1 t1 + fpt t1 + max_length_of_priority_inversion j t1.

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

1 subgoal (ID 2540)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ============================
  t1 <= t1 + fpt <= t1 + max_length_of_priority_inversion j t1

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


          Proof.
            move: (H_jlp_is_scheduled) ⇒ ARRs; apply H_jobs_come_from_arrival_sequence in ARRs.

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

1 subgoal (ID 2543)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  ============================
  t1 <= t1 + fpt <= t1 + max_length_of_priority_inversion j t1

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


            move: H_progr_le_max_nonp_segment ⇒ /andP [GE LE].

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

1 subgoal (ID 2585)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  GE : progr_t1 <= progr_t1 + fpt
  LE : progr_t1 + fpt <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε)
  ============================
  t1 <= t1 + fpt <= t1 + max_length_of_priority_inversion j t1

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


            apply/andP; split; first by rewrite leq_addr.

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

1 subgoal (ID 2612)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  GE : progr_t1 <= progr_t1 + fpt
  LE : progr_t1 + fpt <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε)
  ============================
  t1 + fpt <= t1 + max_length_of_priority_inversion j t1

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


            rewrite leq_add2l.

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

1 subgoal (ID 2621)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  GE : progr_t1 <= progr_t1 + fpt
  LE : progr_t1 + fpt <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε)
  ============================
  fpt <= max_length_of_priority_inversion j t1

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


            unfold max_length_of_priority_inversion.

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

1 subgoal (ID 2623)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  GE : progr_t1 <= progr_t1 + fpt
  LE : progr_t1 + fpt <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε)
  ============================
  fpt <=
  \max_(j_lp <- arrivals_before arr_seq t1 | ~~ hep_job j_lp j)
     (job_max_nonpreemptive_segment j_lp - ε)

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


            rewrite (big_rem jlp) //=.

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

2 subgoals (ID 2640)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  GE : progr_t1 <= progr_t1 + fpt
  LE : progr_t1 + fpt <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε)
  ============================
  fpt <=
  maxn
    (if ~~ hep_job jlp j then job_max_nonpreemptive_segment jlp - ε else 0)
    (\max_(y <- rem (T:=Job) jlp (arrivals_before arr_seq t1) | 
     ~~ hep_job y j) (job_max_nonpreemptive_segment y - ε))

subgoal 2 (ID 2663) is:
 jlp \in arrivals_before arr_seq t1

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


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

1 subgoal (ID 2640)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  GE : progr_t1 <= progr_t1 + fpt
  LE : progr_t1 + fpt <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε)
  ============================
  fpt <=
  maxn
    (if ~~ hep_job jlp j then job_max_nonpreemptive_segment jlp - ε else 0)
    (\max_(y <- rem (T:=Job) jlp (arrivals_before arr_seq t1) | 
     ~~ hep_job y j) (job_max_nonpreemptive_segment y - ε))

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


rewrite H_jlp_low_priority; simpl.

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

1 subgoal (ID 2689)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  GE : progr_t1 <= progr_t1 + fpt
  LE : progr_t1 + fpt <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε)
  ============================
  fpt <=
  maxn (job_max_nonpreemptive_segment jlp - ε)
    (\max_(y <- rem (T:=Job) jlp (arrivals_before arr_seq t1) | 
     ~~ hep_job y j) (job_max_nonpreemptive_segment y - ε))

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


              apply leq_trans with (job_max_nonpreemptive_segment jlp - ε); last by rewrite leq_maxl.

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

1 subgoal (ID 2693)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  GE : progr_t1 <= progr_t1 + fpt
  LE : progr_t1 + fpt <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε)
  ============================
  fpt <= job_max_nonpreemptive_segment jlp - ε

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


                by rewrite leq_add2l in LE.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2663)

subgoal 1 (ID 2663) is:
 jlp \in arrivals_before arr_seq t1

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


}

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

1 subgoal (ID 2663)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  GE : progr_t1 <= progr_t1 + fpt
  LE : progr_t1 + fpt <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε)
  ============================
  jlp \in arrivals_before arr_seq t1

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


            eapply arrived_between_implies_in_arrivals; eauto 2.

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

1 subgoal (ID 2754)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  GE : progr_t1 <= progr_t1 + fpt
  LE : progr_t1 + fpt <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε)
  ============================
  arrived_between jlp 0 t1

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


            apply/andP; split; first by done.

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

1 subgoal (ID 2790)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  GE : progr_t1 <= progr_t1 + fpt
  LE : progr_t1 + fpt <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε)
  ============================
  job_arrival jlp < t1

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


            eapply low_priority_job_arrives_before_busy_interval_prefix with t1; eauto 2.

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

1 subgoal (ID 2791)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  progr_t1 := service sched jlp t1 : nat
  fpt : instant
  H_fpt_is_preemptio_point : job_preemptable jlp (progr_t1 + fpt)
  H_fpt_is_first_preemption_point : forall ρ : nat,
                                    progr_t1 <= ρ <=
                                    progr_t1 +
                                    (job_max_nonpreemptive_segment jlp - ε) ->
                                    job_preemptable jlp ρ ->
                                    service sched jlp t1 + fpt <= ρ
  H_progr_le_max_nonp_segment : progr_t1 <= progr_t1 + fpt <=
                                progr_t1 +
                                (job_max_nonpreemptive_segment jlp - ε)
  ARRs : arrives_in arr_seq jlp
  GE : progr_t1 <= progr_t1 + fpt
  LE : progr_t1 + fpt <= progr_t1 + (job_max_nonpreemptive_segment jlp - ε)
  ============================
  t1 <= t1 < t2

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


              by move: (H_busy_interval_prefix) ⇒ [NEM [QT1 [NQT HPJ]]]; apply/andP; split.

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

No more subgoals.

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


          Qed.

        End FirstPreemptionPointOfjlp.

Next we combine the facts above to conclude the lemma.
        Lemma preemption_time_exists_case3:
           pr_t,
            preemption_time sched pr_t
            t1 pr_t t1 + max_length_of_priority_inversion j t1.

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

1 subgoal (ID 2518)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


        Proof.
          set (service := service sched).

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

1 subgoal (ID 2523)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


          have EX: pt,
              ((service jlp t1) pt (service jlp t1) + (job_max_nonpreemptive_segment jlp - 1)) && job_preemptable jlp pt.

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

2 subgoals (ID 2531)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  ============================
  exists pt : nat,
    (service jlp t1 <= pt <=
     service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
    job_preemptable jlp pt

subgoal 2 (ID 2533) is:
 exists pr_t : instant,
   preemption_time sched pr_t /\
   t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


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

1 subgoal (ID 2531)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  ============================
  exists pt : nat,
    (service jlp t1 <= pt <=
     service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
    job_preemptable jlp pt

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


move: (H_jlp_is_scheduled) ⇒ ARRs; apply H_jobs_come_from_arrival_sequence in ARRs.

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

1 subgoal (ID 2536)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  ARRs : arrives_in arr_seq jlp
  ============================
  exists pt : nat,
    (service jlp t1 <= pt <=
     service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
    job_preemptable jlp pt

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


            move: (proj2 (H_valid_model_with_bounded_nonpreemptive_segments) jlp ARRs) =>[_ EXPP].

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

1 subgoal (ID 2552)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  ARRs : arrives_in arr_seq jlp
  EXPP : nonpreemptive_regions_have_bounded_length jlp
  ============================
  exists pt : nat,
    (service jlp t1 <= pt <=
     service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
    job_preemptable jlp pt

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


            specialize (EXPP (service jlp t1)).

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

1 subgoal (ID 2554)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  ARRs : arrives_in arr_seq jlp
  EXPP : 0 <= service jlp t1 <= job_cost jlp ->
         exists pp : duration,
           service jlp t1 <= pp <=
           service jlp t1 + (job_max_nonpreemptive_segment jlp - ε) /\
           job_preemptable jlp pp
  ============================
  exists pt : nat,
    (service jlp t1 <= pt <=
     service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
    job_preemptable jlp pt

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


            feed EXPP.

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

2 subgoals (ID 2555)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  ARRs : arrives_in arr_seq jlp
  EXPP : 0 <= service jlp t1 <= job_cost jlp ->
         exists pp : duration,
           service jlp t1 <= pp <=
           service jlp t1 + (job_max_nonpreemptive_segment jlp - ε) /\
           job_preemptable jlp pp
  ============================
  0 <= service jlp t1 <= job_cost jlp

subgoal 2 (ID 2560) is:
 exists pt : nat,
   (service jlp t1 <= pt <=
    service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
   job_preemptable jlp pt

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


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

1 subgoal (ID 2555)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  ARRs : arrives_in arr_seq jlp
  EXPP : 0 <= service jlp t1 <= job_cost jlp ->
         exists pp : duration,
           service jlp t1 <= pp <=
           service jlp t1 + (job_max_nonpreemptive_segment jlp - ε) /\
           job_preemptable jlp pp
  ============================
  0 <= service jlp t1 <= job_cost jlp

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


apply/andP; split; first by done.

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

1 subgoal (ID 2587)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  ARRs : arrives_in arr_seq jlp
  EXPP : 0 <= service jlp t1 <= job_cost jlp ->
         exists pp : duration,
           service jlp t1 <= pp <=
           service jlp t1 + (job_max_nonpreemptive_segment jlp - ε) /\
           job_preemptable jlp pp
  ============================
  service jlp t1 <= job_cost jlp

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


              apply service_at_most_cost; first by done.

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

1 subgoal (ID 2593)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  ARRs : arrives_in arr_seq jlp
  EXPP : 0 <= service jlp t1 <= job_cost jlp ->
         exists pp : duration,
           service jlp t1 <= pp <=
           service jlp t1 + (job_max_nonpreemptive_segment jlp - ε) /\
           job_preemptable jlp pp
  ============================
  unit_service_proc_model (processor_state Job)

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


                by apply ideal_proc_model_provides_unit_service.

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

2 subgoals (ID 2560)

subgoal 1 (ID 2560) is:
 exists pt : nat,
   (service jlp t1 <= pt <=
    service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
   job_preemptable jlp pt
subgoal 2 (ID 2533) is:
 exists pr_t : instant,
   preemption_time sched pr_t /\
   t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


            }

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

1 subgoal (ID 2560)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  ARRs : arrives_in arr_seq jlp
  EXPP : exists pp : duration,
           service jlp t1 <= pp <=
           service jlp t1 + (job_max_nonpreemptive_segment jlp - ε) /\
           job_preemptable jlp pp
  ============================
  exists pt : nat,
    (service jlp t1 <= pt <=
     service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
    job_preemptable jlp pt

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


            move: EXPP ⇒ [pt [NEQ PP]].

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

1 subgoal (ID 2617)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  ARRs : arrives_in arr_seq jlp
  pt : duration
  NEQ : service jlp t1 <= pt <=
        service jlp t1 + (job_max_nonpreemptive_segment jlp - ε)
  PP : job_preemptable jlp pt
  ============================
  exists pt0 : nat,
    (service jlp t1 <= pt0 <=
     service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
    job_preemptable jlp pt0

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


             pt; apply/andP; split; by done.

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

1 subgoal (ID 2533)

subgoal 1 (ID 2533) is:
 exists pr_t : instant,
   preemption_time sched pr_t /\
   t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


          }

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

1 subgoal (ID 2533)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  EX : exists pt : nat,
         (service jlp t1 <= pt <=
          service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
         job_preemptable jlp pt
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


          move: (ex_minnP EX) ⇒ [sm_pt /andP [NEQ PP] MIN]; clear EX.

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

1 subgoal (ID 2700)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  sm_pt : nat
  NEQ : service jlp t1 <= sm_pt <=
        service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)
  PP : job_preemptable jlp sm_pt
  MIN : forall n : nat,
        (service jlp t1 <= n <=
         service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
        job_preemptable jlp n -> sm_pt <= n
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


          have Fact: Δ, sm_pt = service jlp t1 + Δ.

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

2 subgoals (ID 2705)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  sm_pt : nat
  NEQ : service jlp t1 <= sm_pt <=
        service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)
  PP : job_preemptable jlp sm_pt
  MIN : forall n : nat,
        (service jlp t1 <= n <=
         service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
        job_preemptable jlp n -> sm_pt <= n
  ============================
  exists Δ : nat, sm_pt = service jlp t1 + Δ

subgoal 2 (ID 2707) is:
 exists pr_t : instant,
   preemption_time sched pr_t /\
   t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


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

1 subgoal (ID 2705)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  sm_pt : nat
  NEQ : service jlp t1 <= sm_pt <=
        service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)
  PP : job_preemptable jlp sm_pt
  MIN : forall n : nat,
        (service jlp t1 <= n <=
         service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
        job_preemptable jlp n -> sm_pt <= n
  ============================
  exists Δ : nat, sm_pt = service jlp t1 + Δ

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


(sm_pt - service jlp t1).

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

1 subgoal (ID 2709)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  sm_pt : nat
  NEQ : service jlp t1 <= sm_pt <=
        service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)
  PP : job_preemptable jlp sm_pt
  MIN : forall n : nat,
        (service jlp t1 <= n <=
         service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
        job_preemptable jlp n -> sm_pt <= n
  ============================
  sm_pt = service jlp t1 + (sm_pt - service jlp t1)

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


            apply/eqP; rewrite eq_sym; apply/eqP; rewrite subnKC //.

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

1 subgoal (ID 2803)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  sm_pt : nat
  NEQ : service jlp t1 <= sm_pt <=
        service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)
  PP : job_preemptable jlp sm_pt
  MIN : forall n : nat,
        (service jlp t1 <= n <=
         service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
        job_preemptable jlp n -> sm_pt <= n
  ============================
  service jlp t1 <= sm_pt

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


              by move: NEQ ⇒ /andP [T _].
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2707)

subgoal 1 (ID 2707) is:
 exists pr_t : instant,
   preemption_time sched pr_t /\
   t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


}

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

1 subgoal (ID 2707)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  sm_pt : nat
  NEQ : service jlp t1 <= sm_pt <=
        service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)
  PP : job_preemptable jlp sm_pt
  MIN : forall n : nat,
        (service jlp t1 <= n <=
         service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
        job_preemptable jlp n -> sm_pt <= n
  Fact : exists Δ : nat, sm_pt = service jlp t1 + Δ
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


          move: Fact ⇒ [Δ EQ]; subst sm_pt; rename Δ into sm_pt.

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

1 subgoal (ID 2891)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  sm_pt : nat
  MIN : forall n : nat,
        (service jlp t1 <= n <=
         service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
        job_preemptable jlp n -> service jlp t1 + sm_pt <= n
  PP : job_preemptable jlp (service jlp t1 + sm_pt)
  NEQ : service jlp t1 <= service jlp t1 + sm_pt <=
        service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


           (t1 + sm_pt); split.

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

2 subgoals (ID 2895)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  sm_pt : nat
  MIN : forall n : nat,
        (service jlp t1 <= n <=
         service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
        job_preemptable jlp n -> service jlp t1 + sm_pt <= n
  PP : job_preemptable jlp (service jlp t1 + sm_pt)
  NEQ : service jlp t1 <= service jlp t1 + sm_pt <=
        service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)
  ============================
  preemption_time sched (t1 + sm_pt)

subgoal 2 (ID 2896) is:
 t1 <= t1 + sm_pt <= t1 + max_length_of_priority_inversion j t1

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


          - apply first_preemption_time.

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

4 subgoals (ID 2897)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  sm_pt : nat
  MIN : forall n : nat,
        (service jlp t1 <= n <=
         service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
        job_preemptable jlp n -> service jlp t1 + sm_pt <= n
  PP : job_preemptable jlp (service jlp t1 + sm_pt)
  NEQ : service jlp t1 <= service jlp t1 + sm_pt <=
        service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)
  ============================
  job_preemptable jlp (service.service sched jlp t1 + sm_pt)

subgoal 2 (ID 2898) is:
 forall ρ : nat,
 service.service sched jlp t1 <= ρ <=
 service.service sched jlp t1 + (job_max_nonpreemptive_segment jlp - ε) ->
 job_preemptable jlp ρ -> service.service sched jlp t1 + sm_pt <= ρ
subgoal 3 (ID 2899) is:
 service.service sched jlp t1 <= service.service sched jlp t1 + sm_pt <=
 service.service sched jlp t1 + (job_max_nonpreemptive_segment jlp - ε)
subgoal 4 (ID 2896) is:
 t1 <= t1 + sm_pt <= t1 + max_length_of_priority_inversion j t1

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


            all: unfold service.service; try done.

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

2 subgoals (ID 2903)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  sm_pt : nat
  MIN : forall n : nat,
        (service jlp t1 <= n <=
         service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
        job_preemptable jlp n -> service jlp t1 + sm_pt <= n
  PP : job_preemptable jlp (service jlp t1 + sm_pt)
  NEQ : service jlp t1 <= service jlp t1 + sm_pt <=
        service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)
  ============================
  forall ρ : nat,
  service_during sched jlp 0 t1 <= ρ <=
  service_during sched jlp 0 t1 + (job_max_nonpreemptive_segment jlp - ε) ->
  job_preemptable jlp ρ -> service_during sched jlp 0 t1 + sm_pt <= ρ

subgoal 2 (ID 2907) is:
 t1 <= t1 + sm_pt <= t1 + max_length_of_priority_inversion j t1

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


            intros; apply MIN; apply/andP; split; by done.

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

1 subgoal (ID 2907)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  sm_pt : nat
  MIN : forall n : nat,
        (service jlp t1 <= n <=
         service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
        job_preemptable jlp n -> service jlp t1 + sm_pt <= n
  PP : job_preemptable jlp (service jlp t1 + sm_pt)
  NEQ : service jlp t1 <= service jlp t1 + sm_pt <=
        service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)
  ============================
  t1 <= t1 + sm_pt <= t1 + max_length_of_priority_inversion j t1

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


          - apply preemption_time_le_max_len_of_priority_inversion.

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

1 subgoal (ID 2989)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  jlp : Job
  H_jlp_is_scheduled : scheduled_at sched jlp t1
  H_jlp_low_priority : ~~ hep_job jlp j
  service := service.service sched : Job -> instant -> nat
  sm_pt : nat
  MIN : forall n : nat,
        (service jlp t1 <= n <=
         service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)) &&
        job_preemptable jlp n -> service jlp t1 + sm_pt <= n
  PP : job_preemptable jlp (service jlp t1 + sm_pt)
  NEQ : service jlp t1 <= service jlp t1 + sm_pt <=
        service jlp t1 + (job_max_nonpreemptive_segment jlp - 1)
  ============================
  service.service sched jlp t1 <= service.service sched jlp t1 + sm_pt <=
  service.service sched jlp t1 + (job_max_nonpreemptive_segment jlp - ε)

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


            unfold service.service; try done.

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

No more subgoals.

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


        Qed.

      End Case3.

    End CaseAnalysis.

By doing the case analysis, we show that indeed there is a preemption time in time interval [[t1, t1 + max_length_of_priority_inversion j t1].
    Lemma preemption_time_exists:
       pr_t,
        preemption_time sched pr_t
        t1 pr_t t1 + max_length_of_priority_inversion j t1.

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

1 subgoal (ID 2513)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

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


    Proof.
      set (service := service sched).

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

1 subgoal (ID 2518)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  service := service.service sched : Job -> instant -> nat
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

----------------------------------------------------------------------------- *)


      move: (H_valid_model_with_bounded_nonpreemptive_segments) ⇒ CORR.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2520)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

----------------------------------------------------------------------------- *)


      move: (H_busy_interval_prefix) ⇒ [NEM [QT1 [NQT HPJ]]].

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2552)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  NEM : t1 < t2
  QT1 : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

----------------------------------------------------------------------------- *)


      ideal_proc_model_sched_case_analysis sched t1 s.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2560)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  NEM : t1 < t2
  QT1 : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  Idle : is_idle sched t1
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

subgoal 2 (ID 2565) is:
 exists pr_t : instant,
   preemption_time sched pr_t /\
   t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

----------------------------------------------------------------------------- *)


      - by apply preemption_time_exists_case1.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2565)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  NEM : t1 < t2
  QT1 : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  s : Job
  Sched_s : scheduled_at sched s t1
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

----------------------------------------------------------------------------- *)


      - destruct (hep_job s j) eqn:PRIO.

(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 2578)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  NEM : t1 < t2
  QT1 : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  s : Job
  Sched_s : scheduled_at sched s t1
  PRIO : hep_job s j = true
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

subgoal 2 (ID 2579) is:
 exists pr_t : instant,
   preemption_time sched pr_t /\
   t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

----------------------------------------------------------------------------- *)


        + by eapply preemption_time_exists_case2; eauto.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2579)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  NEM : t1 < t2
  QT1 : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  s : Job
  Sched_s : scheduled_at sched s t1
  PRIO : hep_job s j = false
  ============================
  exists pr_t : instant,
    preemption_time sched pr_t /\
    t1 <= pr_t <= t1 + max_length_of_priority_inversion j t1

----------------------------------------------------------------------------- *)


        + eapply preemption_time_exists_case3 with s; eauto.

(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 2588)
  
  Task : TaskType
  H : TaskCost Task
  Job : JobType
  H0 : JobTask Job Task
  H1 : JobArrival Job
  H2 : JobCost Job
  arr_seq : arrival_sequence Job
  H_arrival_times_are_consistent : consistent_arrival_times arr_seq
  sched : schedule (processor_state Job)
  H_jobs_come_from_arrival_sequence : jobs_come_from_arrival_sequence sched
                                        arr_seq
  H_jobs_must_arrive_to_execute : jobs_must_arrive_to_execute sched
  H_completed_jobs_dont_execute : completed_jobs_dont_execute sched
  H3 : JLFP_policy Job
  H_priority_is_reflexive : reflexive_priorities
  H_priority_is_transitive : transitive_priorities
  H4 : TaskMaxNonpreemptiveSegment Task
  H5 : JobPreemptable Job
  H_valid_model_with_bounded_nonpreemptive_segments : 
  valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  H_work_conserving : work_conserving arr_seq sched
  H_respects_policy : respects_policy_at_preemption_point arr_seq sched
  job_scheduled_at := scheduled_at sched : Job -> instant -> bool
  job_completed_by := completed_by sched : Job -> instant -> bool
  j : Job
  H_j_arrives : arrives_in arr_seq j
  H_job_cost_positive : job_cost_positive j
  t1, t2 : instant
  H_busy_interval_prefix : busy_interval_prefix arr_seq sched j t1 t2
  service := service.service sched : Job -> instant -> nat
  CORR : valid_model_with_bounded_nonpreemptive_segments arr_seq sched
  NEM : t1 < t2
  QT1 : quiet_time arr_seq sched j t1
  NQT : forall t : nat, t1 < t < t2 -> ~ quiet_time arr_seq sched j t
  HPJ : t1 <= job_arrival j < t2
  s : Job
  Sched_s : scheduled_at sched s t1
  PRIO : hep_job s j = false
  ============================
  ~~ hep_job s j

----------------------------------------------------------------------------- *)


            by rewrite -eqbF_neg; apply /eqP.

(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

----------------------------------------------------------------------------- *)


    Qed.

  End PreemptionTimeExists.

End PriorityInversionIsBounded.