Library prosa.analysis.facts.transform.swaps


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

Welcome to Coq 8.11.2 (June 2020)

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


Require Export prosa.analysis.facts.transform.replace_at.
Require Export prosa.analysis.facts.behavior.deadlines.

In this file, we establish invariants about schedules in which two allocations have been swapped, as for instance it is done in the classic EDF optimality proof.

Section SwappedFacts.
For any given type of jobs...
  Context {Job : JobType}.
... any given type of processor states:
  Context {PState: Type}.
  Context `{ProcessorState Job PState}.

...consider any given reference schedule.
  Variable sched: schedule PState.

Suppose we are given two specific times [t1] and [t2].
  Variable t1 t2: instant.

In the following, let [sched'] denote the schedule in which the allocations at [t1] and [t2] have been swapped.
  Let sched' := swapped sched t1 t2.

First, we note that the trivial case where t1 == t2 is not interesting because then the two schedules are identical.
  Lemma trivial_swap:
    t1 = t2
     t,
      sched t = sched' t.

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

1 subgoal (ID 338)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  ============================
  t1 = t2 -> forall t : instant, sched t = sched' t

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


  Proof.
    rewrite /sched' ⇒ <- t.

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

1 subgoal (ID 344)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  t : instant
  ============================
  sched t = swapped sched t1 t1 t

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


    rewrite /swapped /replace_at.

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

1 subgoal (ID 350)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  t : instant
  ============================
  sched t =
  (if t1 == t then sched t1 else if t1 == t then sched t1 else sched t)

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


    case: (boolP (t1 == t)); last by done.

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

1 subgoal (ID 360)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  t : instant
  ============================
  t1 == t -> sched t = sched t1

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


    by move /eqP →.

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

No more subgoals.

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


  Qed.

In this trivial case, the amount of service received hence is obviously always identical.
  Lemma trivial_swap_service_invariant:
    t1 = t2
     t j,
      service sched j t = service sched' j t.

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

1 subgoal (ID 353)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  ============================
  t1 = t2 ->
  forall (t : instant) (j : Job), service sched j t = service sched' j t

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


  Proof.
    moveSAME t j.

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

1 subgoal (ID 356)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  SAME : t1 = t2
  t : instant
  j : Job
  ============================
  service sched j t = service sched' j t

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


    rewrite /service /service_during /service_at.

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

1 subgoal (ID 377)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  SAME : t1 = t2
  t : instant
  j : Job
  ============================
  \sum_(0 <= t0 < t) service_in j (sched t0) =
  \sum_(0 <= t0 < t) service_in j (sched' t0)

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


    apply eq_big_natt' RANGE.

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

1 subgoal (ID 380)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  SAME : t1 = t2
  t : instant
  j : Job
  t' : nat
  RANGE : 0 <= t' < t
  ============================
  service_in j (sched t') = service_in j (sched' t')

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


    by rewrite trivial_swap.

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

No more subgoals.

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


  Qed.

In any case, the two schedules do not differ at non-swapped times.
  Lemma swap_other_times_invariant:
     t,
      t t1
      t t2
      sched t = sched' t.

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

1 subgoal (ID 362)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  ============================
  forall t : instant, t <> t1 -> t <> t2 -> sched t = sched' t

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


  Proof.
    movet NOT_t1 NOT_t2.

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

1 subgoal (ID 365)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  t : instant
  NOT_t1 : t <> t1
  NOT_t2 : t <> t2
  ============================
  sched t = sched' t

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


    by rewrite /sched' /swapped !rest_of_schedule_invariant //.

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

No more subgoals.

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


  Qed.

By definition, if a job is scheduled at t2 in the original schedule, then it is found at t1 in the new schedule.
  Lemma swap_job_scheduled_t1:
     j,
      scheduled_at sched' j t1 =
      scheduled_at sched j t2.

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

1 subgoal (ID 374)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  ============================
  forall j : Job, scheduled_at sched' j t1 = scheduled_at sched j t2

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


  Proof.
    movej.

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

1 subgoal (ID 375)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  ============================
  scheduled_at sched' j t1 = scheduled_at sched j t2

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


    rewrite /scheduled_at /sched' /swapped /replace_at.

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

1 subgoal (ID 389)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  ============================
  scheduled_in j
    (if t2 == t1 then sched t1 else if t1 == t1 then sched t2 else sched t1) =
  scheduled_in j (sched t2)

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


    case: (boolP(t2 == t1)) ⇒ [/eqP EQ|NE].

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

2 subgoals (ID 435)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  EQ : t2 = t1
  ============================
  scheduled_in j (sched t1) = scheduled_in j (sched t2)

subgoal 2 (ID 434) is:
 scheduled_in j (if t1 == t1 then sched t2 else sched t1) =
 scheduled_in j (sched t2)

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


    - by rewrite EQ.

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

1 subgoal (ID 434)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  NE : t2 != t1
  ============================
  scheduled_in j (if t1 == t1 then sched t2 else sched t1) =
  scheduled_in j (sched t2)

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


    - by rewrite ifT.

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

No more subgoals.

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


  Qed.

Similarly, a job scheduled at t1 in the original schedule is scheduled at t2 after the swap.
  Lemma swap_job_scheduled_t2:
     j,
      scheduled_at sched' j t2 =
      scheduled_at sched j t1.

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

1 subgoal (ID 386)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  ============================
  forall j : Job, scheduled_at sched' j t2 = scheduled_at sched j t1

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


  Proof.
    movej.

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

1 subgoal (ID 387)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  ============================
  scheduled_at sched' j t2 = scheduled_at sched j t1

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


    rewrite /sched' /scheduled_at /swapped /replace_at.

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

1 subgoal (ID 401)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  ============================
  scheduled_in j
    (if t2 == t2 then sched t1 else if t1 == t2 then sched t2 else sched t2) =
  scheduled_in j (sched t1)

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


    case: (boolP(t2 == t1)) ⇒ [/eqP EQ|NE].

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

2 subgoals (ID 447)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  EQ : t2 = t1
  ============================
  scheduled_in j
    (if t2 == t2 then sched t1 else if t1 == t2 then sched t2 else sched t2) =
  scheduled_in j (sched t1)

subgoal 2 (ID 446) is:
 scheduled_in j
   (if t2 == t2 then sched t1 else if t1 == t2 then sched t2 else sched t2) =
 scheduled_in j (sched t1)

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


    - by rewrite ifT.

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

1 subgoal (ID 446)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  NE : t2 != t1
  ============================
  scheduled_in j
    (if t2 == t2 then sched t1 else if t1 == t2 then sched t2 else sched t2) =
  scheduled_in j (sched t1)

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


    - by rewrite ifT.

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

No more subgoals.

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


  Qed.

If a job is scheduled at any time not involved in the swap, then it remains scheduled at that time in the new schedule.
  Lemma swap_job_scheduled_other_times:
     j t,
      t1 != t
      t2 != t
      scheduled_at sched' j t =
      scheduled_at sched j t.

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

1 subgoal (ID 401)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  ============================
  forall (j : Job) (t : nat_eqType),
  t1 != t -> t2 != t -> scheduled_at sched' j t = scheduled_at sched j t

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


  Proof.
    movej t /eqP NOT_t1 /eqP NOT_t2.

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

1 subgoal (ID 488)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : nat_eqType
  NOT_t1 : t1 <> t
  NOT_t2 : t2 <> t
  ============================
  scheduled_at sched' j t = scheduled_at sched j t

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


    rewrite /scheduled_at.

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

1 subgoal (ID 495)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : nat_eqType
  NOT_t1 : t1 <> t
  NOT_t2 : t2 <> t
  ============================
  scheduled_in j (sched' t) = scheduled_in j (sched t)

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


    by rewrite swap_other_times_invariant //; apply: not_eq_sym.

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

No more subgoals.

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


  Qed.

To make case analysis more convenient, we summarize the preceding three lemmas as a disjunction.
  Corollary swap_job_scheduled_cases:
     j t,
      scheduled_at sched' j t
      scheduled_at sched' j t = scheduled_at sched j t
      
      t = t1 scheduled_at sched' j t = scheduled_at sched j t2
      
      t = t2 scheduled_at sched' j t = scheduled_at sched j t1.

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

1 subgoal (ID 442)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  ============================
  forall (j : Job) (t : instant),
  scheduled_at sched' j t ->
  scheduled_at sched' j t = scheduled_at sched j t \/
  t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
  t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

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


  Proof.
    movej t SCHED.

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

1 subgoal (ID 445)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  ============================
  scheduled_at sched' j t = scheduled_at sched j t \/
  t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
  t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

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


    destruct (t1 == t) eqn:T1.

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

2 subgoals (ID 456)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  T1 : (t1 == t) = true
  ============================
  scheduled_at sched' j t = scheduled_at sched j t \/
  t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
  t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

subgoal 2 (ID 457) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
 t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

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


    - right.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 459)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  T1 : (t1 == t) = true
  ============================
  t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
  t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

subgoal 2 (ID 457) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
 t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

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


left.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 461)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  T1 : (t1 == t) = true
  ============================
  t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2

subgoal 2 (ID 457) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
 t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

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


move: T1 ⇒ /eqP T1.

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

2 subgoals (ID 497)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  T1 : t1 = t
  ============================
  t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2

subgoal 2 (ID 457) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
 t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

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


      split ⇒ //.

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

2 subgoals (ID 500)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  T1 : t1 = t
  ============================
  scheduled_at sched' j t = scheduled_at sched j t2

subgoal 2 (ID 457) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
 t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

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


      by rewrite -swap_job_scheduled_t1 T1.

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

1 subgoal (ID 457)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  T1 : (t1 == t) = false
  ============================
  scheduled_at sched' j t = scheduled_at sched j t \/
  t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
  t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

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


    - destruct (t2 == t) eqn:T2.

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

2 subgoals (ID 557)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  T1 : (t1 == t) = false
  T2 : (t2 == t) = true
  ============================
  scheduled_at sched' j t = scheduled_at sched j t \/
  t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
  t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

subgoal 2 (ID 558) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
 t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

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


      + right.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 560)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  T1 : (t1 == t) = false
  T2 : (t2 == t) = true
  ============================
  t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
  t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

subgoal 2 (ID 558) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
 t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

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


right.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 562)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  T1 : (t1 == t) = false
  T2 : (t2 == t) = true
  ============================
  t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

subgoal 2 (ID 558) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
 t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

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


move: T2 ⇒ /eqP T2.

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

2 subgoals (ID 598)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  T1 : (t1 == t) = false
  T2 : t2 = t
  ============================
  t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

subgoal 2 (ID 558) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
 t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

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


        split ⇒ //.

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

2 subgoals (ID 601)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  T1 : (t1 == t) = false
  T2 : t2 = t
  ============================
  scheduled_at sched' j t = scheduled_at sched j t1

subgoal 2 (ID 558) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
 t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

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


        by rewrite -swap_job_scheduled_t2 T2.

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

1 subgoal (ID 558)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  T1 : (t1 == t) = false
  T2 : (t2 == t) = false
  ============================
  scheduled_at sched' j t = scheduled_at sched j t \/
  t = t1 /\ scheduled_at sched' j t = scheduled_at sched j t2 \/
  t = t2 /\ scheduled_at sched' j t = scheduled_at sched j t1

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


      + left.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 649)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  T1 : (t1 == t) = false
  T2 : (t2 == t) = false
  ============================
  scheduled_at sched' j t = scheduled_at sched j t

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


move: T1 T2 ⇒ /eqP/eqP T1 /eqP/eqP T2.

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

1 subgoal (ID 850)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  T1 : t1 != t
  T2 : t2 != t
  ============================
  scheduled_at sched' j t = scheduled_at sched j t

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


        by rewrite -swap_job_scheduled_other_times.

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

No more subgoals.

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


  Qed.

From this, we can easily conclude that no jobs have appeared out of thin air: if a job scheduled at some time in the new schedule, then it was also scheduled at some time in the original schedule.
  Corollary swap_job_scheduled:
     j t,
      scheduled_at sched' j t
       t',
        scheduled_at sched j t'.

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

1 subgoal (ID 455)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  ============================
  forall (j : Job) (t : instant),
  scheduled_at sched' j t -> exists t' : instant, scheduled_at sched j t'

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


  Proof.
    movej t SCHED.

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

1 subgoal (ID 458)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  ============================
  exists t' : instant, scheduled_at sched j t'

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


    move: (SCHED).

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

1 subgoal (ID 459)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  ============================
  scheduled_at sched' j t -> exists t' : instant, scheduled_at sched j t'

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


    move: (swap_job_scheduled_cases j t SCHED) ⇒ [->|[[_ ->] | [_ ->]]].

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

3 subgoals (ID 476)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  ============================
  scheduled_at sched j t -> exists t' : instant, scheduled_at sched j t'

subgoal 2 (ID 517) is:
 scheduled_at sched j t2 -> exists t' : instant, scheduled_at sched j t'
subgoal 3 (ID 518) is:
 scheduled_at sched j t1 -> exists t' : instant, scheduled_at sched j t'

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


    - by t.

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

2 subgoals (ID 517)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  ============================
  scheduled_at sched j t2 -> exists t' : instant, scheduled_at sched j t'

subgoal 2 (ID 518) is:
 scheduled_at sched j t1 -> exists t' : instant, scheduled_at sched j t'

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


    - by t2.

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

1 subgoal (ID 518)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  ============================
  scheduled_at sched j t1 -> exists t' : instant, scheduled_at sched j t'

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


    - by t1.

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

No more subgoals.

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


  Qed.

Mirroring swap_job_scheduled_cases above, we also state a disjunction for case analysis under the premise that a job is scheduled in the original schedule.
  Lemma swap_job_scheduled_original_cases:
     j t,
      scheduled_at sched j t
      scheduled_at sched' j t = scheduled_at sched j t
      
      t = t1 scheduled_at sched' j t2 = scheduled_at sched j t
      
      t = t2 scheduled_at sched' j t1 = scheduled_at sched j t.

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

1 subgoal (ID 496)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  ============================
  forall (j : Job) (t : instant),
  scheduled_at sched j t ->
  scheduled_at sched' j t = scheduled_at sched j t \/
  t = t1 /\ scheduled_at sched' j t2 = scheduled_at sched j t \/
  t = t2 /\ scheduled_at sched' j t1 = scheduled_at sched j t

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


  Proof.
    movej t SCHED.

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

1 subgoal (ID 499)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  ============================
  scheduled_at sched' j t = scheduled_at sched j t \/
  t = t1 /\ scheduled_at sched' j t2 = scheduled_at sched j t \/
  t = t2 /\ scheduled_at sched' j t1 = scheduled_at sched j t

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


    destruct (t1 == t) eqn:T1.

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

2 subgoals (ID 510)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  T1 : (t1 == t) = true
  ============================
  scheduled_at sched' j t = scheduled_at sched j t \/
  t = t1 /\ scheduled_at sched' j t2 = scheduled_at sched j t \/
  t = t2 /\ scheduled_at sched' j t1 = scheduled_at sched j t

subgoal 2 (ID 511) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t2 = scheduled_at sched j t \/
 t = t2 /\ scheduled_at sched' j t1 = scheduled_at sched j t

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


    - right; left.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 515)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  T1 : (t1 == t) = true
  ============================
  t = t1 /\ scheduled_at sched' j t2 = scheduled_at sched j t

subgoal 2 (ID 511) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t2 = scheduled_at sched j t \/
 t = t2 /\ scheduled_at sched' j t1 = scheduled_at sched j t

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


move: T1 ⇒ /eqP T1.

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

2 subgoals (ID 551)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  T1 : t1 = t
  ============================
  t = t1 /\ scheduled_at sched' j t2 = scheduled_at sched j t

subgoal 2 (ID 511) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t2 = scheduled_at sched j t \/
 t = t2 /\ scheduled_at sched' j t1 = scheduled_at sched j t

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


      split ⇒ //.

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

2 subgoals (ID 554)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  T1 : t1 = t
  ============================
  scheduled_at sched' j t2 = scheduled_at sched j t

subgoal 2 (ID 511) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t2 = scheduled_at sched j t \/
 t = t2 /\ scheduled_at sched' j t1 = scheduled_at sched j t

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


      by rewrite swap_job_scheduled_t2 T1.

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

1 subgoal (ID 511)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  T1 : (t1 == t) = false
  ============================
  scheduled_at sched' j t = scheduled_at sched j t \/
  t = t1 /\ scheduled_at sched' j t2 = scheduled_at sched j t \/
  t = t2 /\ scheduled_at sched' j t1 = scheduled_at sched j t

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


    - destruct (t2 == t) eqn:T2.

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

2 subgoals (ID 611)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  T1 : (t1 == t) = false
  T2 : (t2 == t) = true
  ============================
  scheduled_at sched' j t = scheduled_at sched j t \/
  t = t1 /\ scheduled_at sched' j t2 = scheduled_at sched j t \/
  t = t2 /\ scheduled_at sched' j t1 = scheduled_at sched j t

subgoal 2 (ID 612) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t2 = scheduled_at sched j t \/
 t = t2 /\ scheduled_at sched' j t1 = scheduled_at sched j t

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


      + right; right.
(* ----------------------------------[ coqtop ]---------------------------------

2 subgoals (ID 616)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  T1 : (t1 == t) = false
  T2 : (t2 == t) = true
  ============================
  t = t2 /\ scheduled_at sched' j t1 = scheduled_at sched j t

subgoal 2 (ID 612) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t2 = scheduled_at sched j t \/
 t = t2 /\ scheduled_at sched' j t1 = scheduled_at sched j t

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


move: T2 ⇒ /eqP T2.

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

2 subgoals (ID 652)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  T1 : (t1 == t) = false
  T2 : t2 = t
  ============================
  t = t2 /\ scheduled_at sched' j t1 = scheduled_at sched j t

subgoal 2 (ID 612) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t2 = scheduled_at sched j t \/
 t = t2 /\ scheduled_at sched' j t1 = scheduled_at sched j t

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


        split ⇒ //.

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

2 subgoals (ID 655)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  T1 : (t1 == t) = false
  T2 : t2 = t
  ============================
  scheduled_at sched' j t1 = scheduled_at sched j t

subgoal 2 (ID 612) is:
 scheduled_at sched' j t = scheduled_at sched j t \/
 t = t1 /\ scheduled_at sched' j t2 = scheduled_at sched j t \/
 t = t2 /\ scheduled_at sched' j t1 = scheduled_at sched j t

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


        by rewrite swap_job_scheduled_t1 T2.

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

1 subgoal (ID 612)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  T1 : (t1 == t) = false
  T2 : (t2 == t) = false
  ============================
  scheduled_at sched' j t = scheduled_at sched j t \/
  t = t1 /\ scheduled_at sched' j t2 = scheduled_at sched j t \/
  t = t2 /\ scheduled_at sched' j t1 = scheduled_at sched j t

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


      + left.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 703)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  T1 : (t1 == t) = false
  T2 : (t2 == t) = false
  ============================
  scheduled_at sched' j t = scheduled_at sched j t

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


move: T1 T2 ⇒ /eqP/eqP T1 /eqP/eqP T2.

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

1 subgoal (ID 904)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  T1 : t1 != t
  T2 : t2 != t
  ============================
  scheduled_at sched' j t = scheduled_at sched j t

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


        by rewrite swap_job_scheduled_other_times.

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

No more subgoals.

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


  Qed.

Thus, we can conclude that no jobs are lost: if a job is scheduled at some point in the original schedule, then it is also scheduled at some point in the new schedule.
  Corollary swap_job_scheduled_original:
     j t,
      scheduled_at sched j t
       t',
        scheduled_at sched' j t'.

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

1 subgoal (ID 509)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  ============================
  forall (j : Job) (t : instant),
  scheduled_at sched j t -> exists t' : instant, scheduled_at sched' j t'

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


  Proof.
    movej t SCHED.

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

1 subgoal (ID 512)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  ============================
  exists t' : instant, scheduled_at sched' j t'

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


    move: (SCHED).

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

1 subgoal (ID 513)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  ============================
  scheduled_at sched j t -> exists t' : instant, scheduled_at sched' j t'

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


    move: (swap_job_scheduled_original_cases j t SCHED) ⇒ [<-|[[_ <-] | [_ <-]]].

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

3 subgoals (ID 530)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  ============================
  scheduled_at sched' j t -> exists t' : instant, scheduled_at sched' j t'

subgoal 2 (ID 571) is:
 scheduled_at sched' j t2 -> exists t' : instant, scheduled_at sched' j t'
subgoal 3 (ID 572) is:
 scheduled_at sched' j t1 -> exists t' : instant, scheduled_at sched' j t'

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


    - by t.

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

2 subgoals (ID 571)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  ============================
  scheduled_at sched' j t2 -> exists t' : instant, scheduled_at sched' j t'

subgoal 2 (ID 572) is:
 scheduled_at sched' j t1 -> exists t' : instant, scheduled_at sched' j t'

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


    - by t2.

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

1 subgoal (ID 572)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  t : instant
  SCHED : scheduled_at sched j t
  ============================
  scheduled_at sched' j t1 -> exists t' : instant, scheduled_at sched' j t'

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


    - by t1.

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

No more subgoals.

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


  Qed.

In the following, we lift the above basic invariants to statements about invariants about the cumulative amount of service received.
To avoid dealing with symmetric cases, assume in the following that t1 and t2 are ordered.
  Hypothesis H_well_ordered: t1 t2.

As another trivial invariant, we observe that nothing has changed before time t1.
  Lemma swap_before_invariant:
     t,
      t < t1
      sched t = sched' t.

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

1 subgoal (ID 513)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  ============================
  forall t : nat, t < t1 -> sched t = sched' t

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


  Proof.
    movet t_lt_t1.

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

1 subgoal (ID 515)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : nat
  t_lt_t1 : t < t1
  ============================
  sched t = sched' t

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


    move: (ltn_leq_trans t_lt_t1 H_well_ordered) ⇒ t_lt_t2.

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

1 subgoal (ID 520)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : nat
  t_lt_t1 : t < t1
  t_lt_t2 : t < t2
  ============================
  sched t = sched' t

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


    by apply swap_other_times_invariant ⇒ /eqP /eqP EQ; subst;
       [move: t_lt_t1|move: t_lt_t2]; rewrite ltnn.

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

No more subgoals.

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


  Qed.

Similarly, nothing has changed after time t2.
  Lemma swap_after_invariant:
     t,
      t2 < t
      sched t = sched' t.

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

1 subgoal (ID 517)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  ============================
  forall t : nat, t2 < t -> sched t = sched' t

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


  Proof.
    movet t2_lt_t.

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

1 subgoal (ID 519)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : nat
  t2_lt_t : t2 < t
  ============================
  sched t = sched' t

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


    move: (leq_ltn_trans H_well_ordered t2_lt_t) ⇒ t1_lt_t.

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

1 subgoal (ID 524)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : nat
  t2_lt_t : t2 < t
  t1_lt_t : t1 < t
  ============================
  sched t = sched' t

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


    by apply swap_other_times_invariant ⇒ /eqP /eqP EQ; subst;
       [move: t1_lt_t|move: t2_lt_t]; rewrite ltnn.

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

No more subgoals.

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


  Qed.

Thus, we observe that, before t1, the two schedules are identical with regard to the service received by any job because they are identical.
  Corollary service_before_swap_invariant:
     t,
      t t1
       (j : Job),
        service sched j t = service sched' j t.

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

1 subgoal (ID 527)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  ============================
  forall t : nat,
  t <= t1 -> forall j : Job, service sched j t = service sched' j t

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


  Proof.
    movet le_tt1 j.

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

1 subgoal (ID 530)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : nat
  le_tt1 : t <= t1
  j : Job
  ============================
  service sched j t = service sched' j t

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


    rewrite /service.

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

1 subgoal (ID 537)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : nat
  le_tt1 : t <= t1
  j : Job
  ============================
  service_during sched j 0 t = service_during sched' j 0 t

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


    rewrite -!service_at_other_times_invariant //; left; first by done.

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

1 subgoal (ID 652)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : nat
  le_tt1 : t <= t1
  j : Job
  ============================
  t <= t2

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


    apply leq_trans with (n := t1) ⇒ //;
      by apply ltnW.

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

No more subgoals.

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


  Qed.

Likewise, we observe that, *after* t2, the swapped schedule again does not differ with regard to the service received by any job.
  Lemma service_after_swap_invariant:
     t,
      t2 < t
       (j : Job),
        service sched j t = service sched' j t.

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

1 subgoal (ID 537)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  ============================
  forall t : nat,
  t2 < t -> forall j : Job, service sched j t = service sched' j t

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


  Proof.
    movet t2t j.

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

1 subgoal (ID 540)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : nat
  t2t : t2 < t
  j : Job
  ============================
  service sched j t = service sched' j t

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


    move: H_well_ordered.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 541)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : nat
  t2t : t2 < t
  j : Job
  ============================
  t1 <= t2 -> service sched j t = service sched' j t

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


rewrite leq_eqVlt ⇒ /orP [/eqP EQ|t1_lt_t2];
      first by apply trivial_swap_service_invariant.

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

1 subgoal (ID 621)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : nat
  t2t : t2 < t
  j : Job
  t1_lt_t2 : t1 < t2
  ============================
  service sched j t = service sched' j t

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


    have TIME: 0 t1 < t
      by apply /andP; split; try apply ltn_trans with (n := t2); done.

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

1 subgoal (ID 654)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : nat
  t2t : t2 < t
  j : Job
  t1_lt_t2 : t1 < t2
  TIME : 0 <= t1 < t
  ============================
  service sched j t = service sched' j t

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


    rewrite /service !service_in_replaced// /service_at// /replace_at //.

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

1 subgoal (ID 774)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : nat
  t2t : t2 < t
  j : Job
  t1_lt_t2 : t1 < t2
  TIME : 0 <= t1 < t
  ============================
  service_during sched j 0 t =
  service_during sched j 0 t +
  service_in j (if t1 == t1 then sched t2 else sched t1) -
  service_in j (sched t1) +
  service_in j
    (if t2 == t2 then sched t1 else if t1 == t2 then sched t2 else sched t2) -
  service_in j (if t1 == t2 then sched t2 else sched t2)

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


    rewrite ifT// ifT// ifF;
      last by apply ltn_eqF; exact.

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

1 subgoal (ID 863)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : nat
  t2t : t2 < t
  j : Job
  t1_lt_t2 : t1 < t2
  TIME : 0 <= t1 < t
  ============================
  service_during sched j 0 t =
  service_during sched j 0 t + service_in j (sched t2) -
  service_in j (sched t1) + service_in j (sched t1) - 
  service_in j (sched t2)

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


    rewrite subn_abba //.

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

1 subgoal (ID 872)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : nat
  t2t : t2 < t
  j : Job
  t1_lt_t2 : t1 < t2
  TIME : 0 <= t1 < t
  ============================
  service_in j (sched t1) <= service_during sched j 0 t

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


    rewrite -(service_split_at_point _ _ _ t1 _) /service_at //.

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

1 subgoal (ID 927)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : nat
  t2t : t2 < t
  j : Job
  t1_lt_t2 : t1 < t2
  TIME : 0 <= t1 < t
  ============================
  service_in j (sched t1) <=
  service_during sched j 0 t1 + service_in j (sched t1) +
  service_during sched j (succn t1) t

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


    apply leq_trans with (n := service_during sched j 0 t1 + service_in j (sched t1));
      [rewrite addnC|]; by apply leq_addr.

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

No more subgoals.

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


  Qed.

Finally, we note that, trivially, jobs that are not involved in the swap receive invariant service.
  Lemma service_of_others_invariant:
     t j,
      ~~ scheduled_in j (sched t1)
      ~~ scheduled_in j (sched t2)
      service sched j t = service sched' j t.

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

1 subgoal (ID 558)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  ============================
  forall (t : instant) (j : Job),
  ~~ scheduled_in j (sched t1) ->
  ~~ scheduled_in j (sched t2) -> service sched j t = service sched' j t

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


  Proof.
    movet j NOT1 NOT2.

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

1 subgoal (ID 562)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : instant
  j : Job
  NOT1 : ~~ scheduled_in j (sched t1)
  NOT2 : ~~ scheduled_in j (sched t2)
  ============================
  service sched j t = service sched' j t

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


    move: H_well_ordered.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 563)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : instant
  j : Job
  NOT1 : ~~ scheduled_in j (sched t1)
  NOT2 : ~~ scheduled_in j (sched t2)
  ============================
  t1 <= t2 -> service sched j t = service sched' j t

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


rewrite leq_eqVlt ⇒ /orP [/eqP EQ|t1_lt_t2];
      first by apply trivial_swap_service_invariant.

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

1 subgoal (ID 643)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : instant
  j : Job
  NOT1 : ~~ scheduled_in j (sched t1)
  NOT2 : ~~ scheduled_in j (sched t2)
  t1_lt_t2 : t1 < t2
  ============================
  service sched j t = service sched' j t

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


    rewrite /service.

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

1 subgoal (ID 651)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : instant
  j : Job
  NOT1 : ~~ scheduled_in j (sched t1)
  NOT2 : ~~ scheduled_in j (sched t2)
  t1_lt_t2 : t1 < t2
  ============================
  service_during sched j 0 t = service_during sched' j 0 t

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


    rewrite -!service_during_of_others_invariant // !/replace_at //;
            [rewrite ifT// | rewrite ifT// | rewrite ifF// ].

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

1 subgoal (ID 910)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  sched' := swapped sched t1 t2 : schedule PState
  H_well_ordered : t1 <= t2
  t : instant
  j : Job
  NOT1 : ~~ scheduled_in j (sched t1)
  NOT2 : ~~ scheduled_in j (sched t2)
  t1_lt_t2 : t1 < t2
  ============================
  (t1 == t2) = false

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


      by apply ltn_eqF.

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

No more subgoals.

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


  Qed.

End SwappedFacts.

In the following section, we observe that, if the original schedule satisfies certain properties, then so does the new scheduled obtained by swapping the allocations at times [t1] and [t2].
For any given type of jobs...
  Context {Job : JobType} `{JobCost Job} `{JobDeadline Job} `{JobArrival Job}.
... any given type of processor states:
  Context {PState: eqType}.
  Context `{ProcessorState Job PState}.

...consider any given reference schedule.
  Variable sched: schedule PState.

Suppose we are given two specific times [t1] and [t2]...
  Variable t1 t2: instant.

...such that [t1] is no later than [t2].
  Hypothesis H_order: t1 t2.

We let [sched'] denote the schedule in which the allocations at [t1] and [t2] have been swapped.
  Let sched' := swapped sched t1 t2.

First, we observe that if jobs never accumulate more service than required, then that's still the case after the swap.
  Lemma swapped_service_bound:
    ( j t, service sched j t job_cost j)
    ( j t, service sched' j t job_cost j).

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

1 subgoal (ID 351)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  ============================
  (forall (j : Job) (t : instant), service sched j t <= job_cost j) ->
  forall (j : Job) (t : instant), service sched' j t <= job_cost j

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


  Proof.
    moveCOMP.

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

1 subgoal (ID 352)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  COMP : forall (j : Job) (t : instant), service sched j t <= job_cost j
  ============================
  forall (j : Job) (t : instant), service sched' j t <= job_cost j

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


    rewrite /completed_jobs_dont_executej t.

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

1 subgoal (ID 359)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  COMP : forall (j : Job) (t : instant), service sched j t <= job_cost j
  j : Job
  t : instant
  ============================
  service sched' j t <= job_cost j

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


    wlog: t / (t2 < t); last first.

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

2 subgoals (ID 364)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  COMP : forall (j : Job) (t : instant), service sched j t <= job_cost j
  j : Job
  t : instant
  ============================
  t2 < t -> service sched' j t <= job_cost j

subgoal 2 (ID 361) is:
 (forall t0 : instant, t2 < t0 -> service sched' j t0 <= job_cost j) ->
 service sched' j t <= job_cost j

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


    {

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

1 subgoal (ID 364)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  COMP : forall (j : Job) (t : instant), service sched j t <= job_cost j
  j : Job
  t : instant
  ============================
  t2 < t -> service sched' j t <= job_cost j

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


      moveLATE.

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

1 subgoal (ID 365)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  COMP : forall (j : Job) (t : instant), service sched j t <= job_cost j
  j : Job
  t : instant
  LATE : t2 < t
  ============================
  service sched' j t <= job_cost j

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


      move: H_order.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 366)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  COMP : forall (j : Job) (t : instant), service sched j t <= job_cost j
  j : Job
  t : instant
  LATE : t2 < t
  ============================
  t1 <= t2 -> service sched' j t <= job_cost j

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


rewrite leq_eqVlt ⇒ /orP [/eqP EQ| LT].

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

2 subgoals (ID 445)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  COMP : forall (j : Job) (t : instant), service sched j t <= job_cost j
  j : Job
  t : instant
  LATE : t2 < t
  EQ : t1 = t2
  ============================
  service sched' j t <= job_cost j

subgoal 2 (ID 446) is:
 service sched' j t <= job_cost j

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


      - by rewrite -(trivial_swap_service_invariant sched t1 t2) //.

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

1 subgoal (ID 446)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  COMP : forall (j : Job) (t : instant), service sched j t <= job_cost j
  j : Job
  t : instant
  LATE : t2 < t
  LT : t1 < t2
  ============================
  service sched' j t <= job_cost j

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


      - by rewrite -(service_after_swap_invariant sched t1 t2) //.

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

1 subgoal (ID 361)

subgoal 1 (ID 361) is:
 (forall t0 : instant, t2 < t0 -> service sched' j t0 <= job_cost j) ->
 service sched' j t <= job_cost j

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


    }

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

1 subgoal (ID 361)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  COMP : forall (j : Job) (t : instant), service sched j t <= job_cost j
  j : Job
  t : instant
  ============================
  (forall t0 : instant, t2 < t0 -> service sched' j t0 <= job_cost j) ->
  service sched' j t <= job_cost j

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


    {

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

1 subgoal (ID 361)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  COMP : forall (j : Job) (t : instant), service sched j t <= job_cost j
  j : Job
  t : instant
  ============================
  (forall t0 : instant, t2 < t0 -> service sched' j t0 <= job_cost j) ->
  service sched' j t <= job_cost j

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


      moveARGUMENT.

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

1 subgoal (ID 473)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  COMP : forall (j : Job) (t : instant), service sched j t <= job_cost j
  j : Job
  t : instant
  ARGUMENT : forall t : instant, t2 < t -> service sched' j t <= job_cost j
  ============================
  service sched' j t <= job_cost j

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


      case: (boolP (t2 < t)); first by apply (ARGUMENT t).

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

1 subgoal (ID 483)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  COMP : forall (j : Job) (t : instant), service sched j t <= job_cost j
  j : Job
  t : instant
  ARGUMENT : forall t : instant, t2 < t -> service sched' j t <= job_cost j
  ============================
  ~~ (t2 < t) -> service sched' j t <= job_cost j

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


      rewrite -leqNgtEARLY.

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

1 subgoal (ID 488)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  COMP : forall (j : Job) (t : instant), service sched j t <= job_cost j
  j : Job
  t : instant
  ARGUMENT : forall t : instant, t2 < t -> service sched' j t <= job_cost j
  EARLY : t <= t2
  ============================
  service sched' j t <= job_cost j

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


      apply leq_trans with (n := service sched' j t2.+1).

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

2 subgoals (ID 492)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  COMP : forall (j : Job) (t : instant), service sched j t <= job_cost j
  j : Job
  t : instant
  ARGUMENT : forall t : instant, t2 < t -> service sched' j t <= job_cost j
  EARLY : t <= t2
  ============================
  service sched' j t <= service sched' j (succn t2)

subgoal 2 (ID 493) is:
 service sched' j (succn t2) <= job_cost j

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


      - apply service_monotonic.

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

2 subgoals (ID 497)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  COMP : forall (j : Job) (t : instant), service sched j t <= job_cost j
  j : Job
  t : instant
  ARGUMENT : forall t : instant, t2 < t -> service sched' j t <= job_cost j
  EARLY : t <= t2
  ============================
  t <= succn t2

subgoal 2 (ID 493) is:
 service sched' j (succn t2) <= job_cost j

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


        by apply leq_trans with (n := t2) ⇒ //.

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

1 subgoal (ID 493)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  COMP : forall (j : Job) (t : instant), service sched j t <= job_cost j
  j : Job
  t : instant
  ARGUMENT : forall t : instant, t2 < t -> service sched' j t <= job_cost j
  EARLY : t <= t2
  ============================
  service sched' j (succn t2) <= job_cost j

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


      - by apply (ARGUMENT t2.+1).

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

No more subgoals.

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


    }

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

No more subgoals.

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


  Qed.

From the above service bound, we conclude that, if if completed jobs don't execute in the original schedule, then that's still the case after the swap, assuming an ideal unit-service model (i.e., scheduled jobs receive exactly one unit of service).
  Lemma swapped_completed_jobs_dont_execute:
    unit_service_proc_model PState
    ideal_progress_proc_model PState
    completed_jobs_dont_execute sched
    completed_jobs_dont_execute sched'.

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

1 subgoal (ID 364)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  ============================
  unit_service_proc_model PState ->
  ideal_progress_proc_model PState ->
  completed_jobs_dont_execute sched -> completed_jobs_dont_execute sched'

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


  Proof.
    moveUNIT IDEAL COMP.

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

1 subgoal (ID 367)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  UNIT : unit_service_proc_model PState
  IDEAL : ideal_progress_proc_model PState
  COMP : completed_jobs_dont_execute sched
  ============================
  completed_jobs_dont_execute sched'

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


    apply ideal_progress_completed_jobs ⇒ //.

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

1 subgoal (ID 371)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  UNIT : unit_service_proc_model PState
  IDEAL : ideal_progress_proc_model PState
  COMP : completed_jobs_dont_execute sched
  ============================
  forall (j : Job) (t : instant), service sched' j t <= job_cost j

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


    apply swapped_service_bound.

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

1 subgoal (ID 398)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  UNIT : unit_service_proc_model PState
  IDEAL : ideal_progress_proc_model PState
  COMP : completed_jobs_dont_execute sched
  ============================
  forall (j : Job) (t : instant), service sched j t <= job_cost j

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


    by movej; apply service_at_most_cost.

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

No more subgoals.

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


  Qed.

Suppose all jobs in the original schedule come from some arrival sequence...
...then that's still the case in the new schedule.
  Lemma swapped_jobs_come_from_arrival_sequence:
    jobs_come_from_arrival_sequence sched' arr_seq.

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

1 subgoal (ID 371)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  arr_seq : arrival_sequence Job
  H_from_arr_seq : jobs_come_from_arrival_sequence sched arr_seq
  ============================
  jobs_come_from_arrival_sequence sched' arr_seq

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


  Proof.
    move: H_from_arr_seq.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 372)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  arr_seq : arrival_sequence Job
  H_from_arr_seq : jobs_come_from_arrival_sequence sched arr_seq
  ============================
  jobs_come_from_arrival_sequence sched arr_seq ->
  jobs_come_from_arrival_sequence sched' arr_seq

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


rewrite /jobs_come_from_arrival_sequenceARR_sched j t SCHED.

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

1 subgoal (ID 383)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  arr_seq : arrival_sequence Job
  H_from_arr_seq : jobs_come_from_arrival_sequence sched arr_seq
  ARR_sched : forall (j : Job) (t : instant),
              scheduled_at sched j t -> arrives_in arr_seq j
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  ============================
  arrives_in arr_seq j

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


    move: (swap_job_scheduled sched _ _ j t SCHED) ⇒ [t' SCHED'].

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

1 subgoal (ID 402)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  H1 : JobArrival Job
  PState : eqType
  H2 : ProcessorState Job PState
  sched : schedule PState
  t1, t2 : instant
  H_order : t1 <= t2
  sched' := swapped sched t1 t2 : schedule PState
  arr_seq : arrival_sequence Job
  H_from_arr_seq : jobs_come_from_arrival_sequence sched arr_seq
  ARR_sched : forall (j : Job) (t : instant),
              scheduled_at sched j t -> arrives_in arr_seq j
  j : Job
  t : instant
  SCHED : scheduled_at sched' j t
  t' : instant
  SCHED' : scheduled_at sched j t'
  ============================
  arrives_in arr_seq j

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


    by apply (ARR_sched j t' SCHED').

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

No more subgoals.

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


  Qed.

End SwappedScheduleProperties.

In the next section, we consider a special case of swaps: namely, when the job moved earlier has an earlier deadline than the job being moved to a later allocation, assuming that the former has not yet missed its deadline, which is the core transformation of the classic EDF optimality proof.
Section EDFSwap.
For any given type of jobs with costs and deadlines...
  Context {Job : JobType} `{JobCost Job} `{JobDeadline Job}.
... any given type of processor states...
  Context {PState: eqType}.
  Context `{ProcessorState Job PState}.

...consider a given reference schedule...
  Variable sched: schedule PState.

...in which complete jobs don't execute...
...and scheduled jobs always receive service.
Suppose we are given two specific times [t1] and [t2]...
  Variable t1 t2: instant.

...that are ordered.
  Hypothesis H_well_ordered: t1 t2.

Further, assume that, if there are jobs scheduled at times t1 and t2, then they either have the same deadline or violate EDF, ...
  Hypothesis H_not_EDF:
     j1 j2,
      scheduled_at sched j1 t1
      scheduled_at sched j2 t2
      job_deadline j1 job_deadline j2.

...and that we don't move idle times or deadline misses earlier, i.e., if t1 is not an idle time, then neither is t2 and whatever job is scheduled at time t2 has not yet missed its deadline.
  Hypothesis H_no_idle_time_at_t2:
     j1,
      scheduled_at sched j1 t1
       j2, scheduled_at sched j2 t2 job_deadline j2 > t2.

Consider the schedule obtained from swapping the allocations at times t1 and t2.
  Let sched' := swapped sched t1 t2.

The key property of this transformation is that, for any job that meets its deadline in the original schedule, we have not introduced any deadline misses, which we establish by considering a number of different cases.
Consider any job...
    Variable j: Job.

... that meets its deadline in the original schedule.
    Hypothesis H_deadline_met: job_meets_deadline sched j.

First we observe that jobs that are not involved in the swap still meet their deadlines.
    Lemma uninvolved_implies_deadline_met:
      ~~ scheduled_at sched j t1
      ~~ scheduled_at sched j t2
      job_meets_deadline sched' j.

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

1 subgoal (ID 385)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  H_deadline_met : job_meets_deadline sched j
  ============================
  ~~ scheduled_at sched j t1 ->
  ~~ scheduled_at sched j t2 -> job_meets_deadline sched' j

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


    Proof.
      rewrite /scheduled_atNOT_t1 NOT_t2.

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

1 subgoal (ID 394)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  H_deadline_met : job_meets_deadline sched j
  NOT_t1 : ~~ scheduled_in j (sched t1)
  NOT_t2 : ~~ scheduled_in j (sched t2)
  ============================
  job_meets_deadline sched' j

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


      rewrite -service_invariant_implies_deadline_met; first by exact H_deadline_met.

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

1 subgoal (ID 424)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  H_deadline_met : job_meets_deadline sched j
  NOT_t1 : ~~ scheduled_in j (sched t1)
  NOT_t2 : ~~ scheduled_in j (sched t2)
  ============================
  service sched j (job_deadline j) = service sched' j (job_deadline j)

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


      by apply: service_of_others_invariant.

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

No more subgoals.

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


    Qed.

Next, we observe that a swap is unproblematic for the job scheduled at time t2.
    Lemma moved_earlier_implies_deadline_met:
      scheduled_at sched j t2
      job_meets_deadline sched' j.

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

1 subgoal (ID 394)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  H_deadline_met : job_meets_deadline sched j
  ============================
  scheduled_at sched j t2 -> job_meets_deadline sched' j

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


    Proof.
      moveAT_t2.

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

1 subgoal (ID 395)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  H_deadline_met : job_meets_deadline sched j
  AT_t2 : scheduled_at sched j t2
  ============================
  job_meets_deadline sched' j

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


      rewrite -service_invariant_implies_deadline_met; first by exact H_deadline_met.

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

1 subgoal (ID 425)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  H_deadline_met : job_meets_deadline sched j
  AT_t2 : scheduled_at sched j t2
  ============================
  service sched j (job_deadline j) = service sched' j (job_deadline j)

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


      apply service_after_swap_invariant ⇒ //.

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

1 subgoal (ID 430)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  H_deadline_met : job_meets_deadline sched j
  AT_t2 : scheduled_at sched j t2
  ============================
  t2 < job_deadline j

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


      by apply scheduled_at_implies_later_deadline with (sched0 := sched) ⇒ //.

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

No more subgoals.

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


    Qed.

Finally, we observe is also unproblematic for the job that is moved to a later allocation.
    Lemma moved_later_implies_deadline_met:
      scheduled_at sched j t1
      job_meets_deadline sched' j.

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

1 subgoal (ID 403)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  H_deadline_met : job_meets_deadline sched j
  ============================
  scheduled_at sched j t1 -> job_meets_deadline sched' j

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


    Proof.
      moveAT_t1.

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

1 subgoal (ID 404)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  H_deadline_met : job_meets_deadline sched j
  AT_t1 : scheduled_at sched j t1
  ============================
  job_meets_deadline sched' j

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


      move: (H_no_idle_time_at_t2 j AT_t1) ⇒ [j2 [AT_t2 DL2]].

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

1 subgoal (ID 426)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  H_deadline_met : job_meets_deadline sched j
  AT_t1 : scheduled_at sched j t1
  j2 : Job
  AT_t2 : scheduled_at sched j2 t2
  DL2 : t2 < job_deadline j2
  ============================
  job_meets_deadline sched' j

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


      move: (H_not_EDF j j2 AT_t1 AT_t2) ⇒ DL2_le_DL1.

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

1 subgoal (ID 428)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  H_deadline_met : job_meets_deadline sched j
  AT_t1 : scheduled_at sched j t1
  j2 : Job
  AT_t2 : scheduled_at sched j2 t2
  DL2 : t2 < job_deadline j2
  DL2_le_DL1 : job_deadline j2 <= job_deadline j
  ============================
  job_meets_deadline sched' j

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


      rewrite -service_invariant_implies_deadline_met; first by exact H_deadline_met.

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

1 subgoal (ID 458)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  H_deadline_met : job_meets_deadline sched j
  AT_t1 : scheduled_at sched j t1
  j2 : Job
  AT_t2 : scheduled_at sched j2 t2
  DL2 : t2 < job_deadline j2
  DL2_le_DL1 : job_deadline j2 <= job_deadline j
  ============================
  service sched j (job_deadline j) = service sched' j (job_deadline j)

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


      apply service_after_swap_invariant ⇒ //.

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

1 subgoal (ID 463)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  H_deadline_met : job_meets_deadline sched j
  AT_t1 : scheduled_at sched j t1
  j2 : Job
  AT_t2 : scheduled_at sched j2 t2
  DL2 : t2 < job_deadline j2
  DL2_le_DL1 : job_deadline j2 <= job_deadline j
  ============================
  t2 < job_deadline j

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


      apply ltn_leq_trans with (n := job_deadline j2) ⇒ //.

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

No more subgoals.

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


    Qed.

  End NoNewDeadlineMissesCases.

From the above case analysis, we conclude that no deadline misses are introduced in the schedule obtained from swapping the allocations at times t1 and t2.
  Theorem edf_swap_no_deadline_misses_introduced:
     j, job_meets_deadline sched j job_meets_deadline sched' j.

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

1 subgoal (ID 382)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  ============================
  forall j : Job, job_meets_deadline sched j -> job_meets_deadline sched' j

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


  Proof.
    movej DL_MET.

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

1 subgoal (ID 384)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  DL_MET : job_meets_deadline sched j
  ============================
  job_meets_deadline sched' j

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


    case: (boolP (scheduled_at sched j t1)) ⇒ AT_t1.

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

2 subgoals (ID 405)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  DL_MET : job_meets_deadline sched j
  AT_t1 : scheduled_at sched j t1
  ============================
  job_meets_deadline sched' j

subgoal 2 (ID 406) is:
 job_meets_deadline sched' j

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


    - by apply moved_later_implies_deadline_met.

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

1 subgoal (ID 406)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  DL_MET : job_meets_deadline sched j
  AT_t1 : ~~ scheduled_at sched j t1
  ============================
  job_meets_deadline sched' j

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


    - case (boolP (scheduled_at sched j t2)) ⇒ AT_t2.

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

2 subgoals (ID 417)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  DL_MET : job_meets_deadline sched j
  AT_t1 : ~~ scheduled_at sched j t1
  AT_t2 : scheduled_at sched j t2
  ============================
  job_meets_deadline sched' j

subgoal 2 (ID 418) is:
 job_meets_deadline sched' j

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


      + by apply moved_earlier_implies_deadline_met.

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

1 subgoal (ID 418)
  
  Job : JobType
  H : JobCost Job
  H0 : JobDeadline Job
  PState : eqType
  H1 : ProcessorState Job PState
  sched : schedule PState
  H_completed_jobs : completed_jobs_dont_execute sched
  H_scheduled_implies_serviced : ideal_progress_proc_model PState
  t1, t2 : instant
  H_well_ordered : t1 <= t2
  H_not_EDF : forall j1 j2 : Job,
              scheduled_at sched j1 t1 ->
              scheduled_at sched j2 t2 -> job_deadline j2 <= job_deadline j1
  H_no_idle_time_at_t2 : forall j1 : Job,
                         scheduled_at sched j1 t1 ->
                         exists j2 : Job,
                           scheduled_at sched j2 t2 /\ t2 < job_deadline j2
  sched' := swapped sched t1 t2 : schedule PState
  j : Job
  DL_MET : job_meets_deadline sched j
  AT_t1 : ~~ scheduled_at sched j t1
  AT_t2 : ~~ scheduled_at sched j t2
  ============================
  job_meets_deadline sched' j

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


      + by apply uninvolved_implies_deadline_met.

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

No more subgoals.

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


  Qed.

 End EDFSwap.