Library prosa.analysis.facts.transform.replace_at


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

Welcome to Coq 8.11.2 (June 2020)

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


Require Export prosa.analysis.transform.swap.
Require Export prosa.analysis.facts.behavior.completion.

In this file, we make a few simple observations about schedules with replacements.
Section ReplaceAtFacts.

For any given type of jobs...
  Context {Job : JobType}.

... and 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 a specific time [t'] ...
  Variable t': instant.

...and a replacement state [state].
  Variable nstate: PState.

In the following, let [sched'] denote the schedule with replacement at time t'.
We begin with the trivial observation that [replace_at sched t' nstate] indeed returns [nstate] at time [t'].
  Lemma replace_at_def:
    sched' t' = nstate.

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

1 subgoal (ID 331)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  ============================
  sched' t' = nstate

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


  Proof.
    rewrite /sched' /replace_at.

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

1 subgoal (ID 335)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  ============================
  (if t' == t' then nstate else sched t') = nstate

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


    now apply ifT.

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

No more subgoals.

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


  Qed.

Equally trivially, the schedule doesn't change at other times.
  Lemma rest_of_schedule_invariant:
     t, t t' sched' t = sched t.

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

1 subgoal (ID 338)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  ============================
  forall t : instant, t <> t' -> sched' t = sched t

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


  Proof.
    movet DIFF.

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

1 subgoal (ID 340)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  t : instant
  DIFF : t <> t'
  ============================
  sched' t = sched t

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


    rewrite /sched' /replace_at.

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

1 subgoal (ID 344)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  t : instant
  DIFF : t <> t'
  ============================
  (if t' == t then nstate else sched t) = sched t

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


    case (t' == t) eqn:TT; last by done.

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

1 subgoal (ID 357)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  t : instant
  DIFF : t <> t'
  TT : (t' == t) = true
  ============================
  nstate = sched t

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


    by move/eqP in TT; rewrite TT in DIFF; contradiction.

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

No more subgoals.

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


  Qed.

As a result, the service in intervals that do not intersect with [t'] is invariant, too.
  Lemma service_at_other_times_invariant:
     t1 t2,
      t2 t' t' < t1
       j,
        service_during sched j t1 t2
        =
        service_during sched' j t1 t2.

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

1 subgoal (ID 352)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  ============================
  forall t1 t2 : nat,
  t2 <= t' \/ t' < t1 ->
  forall j : Job,
  service_during sched j t1 t2 = service_during sched' j t1 t2

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


  Proof.
    movet1 t2 SWAP_EXCLUDED j.

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

1 subgoal (ID 356)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  t1, t2 : nat
  SWAP_EXCLUDED : t2 <= t' \/ t' < t1
  j : Job
  ============================
  service_during sched j t1 t2 = service_during sched' j t1 t2

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


    rewrite /service_during /service_at.

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

1 subgoal (ID 370)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  t1, t2 : nat
  SWAP_EXCLUDED : t2 <= t' \/ t' < t1
  j : Job
  ============================
  \sum_(t1 <= t < t2) service_in j (sched t) =
  \sum_(t1 <= t < t2) service_in j (sched' t)

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


    apply eq_big_natt /andP [le_t1t lt_tt2].

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

1 subgoal (ID 412)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  t1, t2 : nat
  SWAP_EXCLUDED : t2 <= t' \/ t' < t1
  j : Job
  t : nat
  le_t1t : t1 <= t
  lt_tt2 : t < t2
  ============================
  service_in j (sched t) = service_in j (sched' t)

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


    rewrite rest_of_schedule_invariant//.

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

1 subgoal (ID 417)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  t1, t2 : nat
  SWAP_EXCLUDED : t2 <= t' \/ t' < t1
  j : Job
  t : nat
  le_t1t : t1 <= t
  lt_tt2 : t < t2
  ============================
  t <> t'

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


    eapply point_not_in_interval; eauto.

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

1 subgoal (ID 448)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  t1, t2 : nat
  SWAP_EXCLUDED : t2 <= t' \/ t' < t1
  j : Job
  t : nat
  le_t1t : t1 <= t
  lt_tt2 : t < t2
  ============================
  t1 <= t < t2

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


    apply /andP; by split.

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

No more subgoals.

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


  Qed.

Next, we consider the amount of service received in intervals that do span across the replacement point. We can relate the service received in the original and the new schedules by adding the service in the respective "missing" state...
  Lemma service_delta:
     t1 t2,
      t1 t' < t2
       j,
        service_during sched j t1 t2 + service_at sched' j t'
        =
        service_during sched' j t1 t2 + service_at sched j t'.

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

1 subgoal (ID 374)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  ============================
  forall t1 t2 : nat,
  t1 <= t' < t2 ->
  forall j : Job,
  service_during sched j t1 t2 + service_at sched' j t' =
  service_during sched' j t1 t2 + service_at sched j t'

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


  Proof.
    movet1 t2 TIMES j.

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

1 subgoal (ID 378)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  t1, t2 : nat
  TIMES : t1 <= t' < t2
  j : Job
  ============================
  service_during sched j t1 t2 + service_at sched' j t' =
  service_during sched' j t1 t2 + service_at sched j t'

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


    rewrite -(service_split_at_point sched _ _ t' _) //
            -(service_split_at_point sched' _ _ t' _) //.

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

1 subgoal (ID 443)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  t1, t2 : nat
  TIMES : t1 <= t' < t2
  j : Job
  ============================
  service_during sched j t1 t' + service_at sched j t' +
  service_during sched j (succn t') t2 + service_at sched' j t' =
  service_during sched' j t1 t' + service_at sched' j t' +
  service_during sched' j (succn t') t2 + service_at sched j t'

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


    by rewrite !service_at_other_times_invariant -/sched'; [ring | right | left].

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

No more subgoals.

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


  Qed.

...which we can also rewrite as follows.
  Corollary service_in_replaced:
     t1 t2,
      t1 t' < t2
       j,
        service_during sched' j t1 t2
        =
        service_during sched j t1 t2 + service_at sched' j t' - service_at sched j t'.

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

1 subgoal (ID 396)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  ============================
  forall t1 t2 : nat,
  t1 <= t' < t2 ->
  forall j : Job,
  service_during sched' j t1 t2 =
  service_during sched j t1 t2 + service_at sched' j t' -
  service_at sched j t'

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


  Proof. movet1 t2 ORDER j.
(* ----------------------------------[ coqtop ]---------------------------------

1 subgoal (ID 400)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  t1, t2 : nat
  ORDER : t1 <= t' < t2
  j : Job
  ============================
  service_during sched' j t1 t2 =
  service_during sched j t1 t2 + service_at sched' j t' -
  service_at sched j t'

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


by rewrite service_delta// addnK.
(* ----------------------------------[ coqtop ]---------------------------------

No more subgoals.

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


Qed.

As another simple invariant (useful for case analysis), we observe that if a job is scheduled neither in the replaced nor in the new state, then at any time it receives exactly the same amount of service in the new schedule with replacements as in the original schedule.
  Lemma service_at_of_others_invariant (j: Job):
    ~~ scheduled_in j (sched' t')
    ~~ scheduled_in j (sched t')
     t,
      service_at sched j t = service_at sched' j t.

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

1 subgoal (ID 413)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  j : Job
  ============================
  ~~ scheduled_in j (sched' t') ->
  ~~ scheduled_in j (sched t') ->
  forall t : instant, service_at sched j t = service_at sched' j t

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


  Proof.
    moveNOT_IN_NEW NOT_IN_OLD t.

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

1 subgoal (ID 416)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  j : Job
  NOT_IN_NEW : ~~ scheduled_in j (sched' t')
  NOT_IN_OLD : ~~ scheduled_in j (sched t')
  t : instant
  ============================
  service_at sched j t = service_at sched' j t

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


    rewrite /service_at.

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

1 subgoal (ID 423)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  j : Job
  NOT_IN_NEW : ~~ scheduled_in j (sched' t')
  NOT_IN_OLD : ~~ scheduled_in j (sched t')
  t : instant
  ============================
  service_in j (sched t) = service_in j (sched' t)

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


    case: (boolP (t == t')) ⇒ /eqP TT.

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

2 subgoals (ID 510)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  j : Job
  NOT_IN_NEW : ~~ scheduled_in j (sched' t')
  NOT_IN_OLD : ~~ scheduled_in j (sched t')
  t : instant
  TT : t = t'
  ============================
  service_in j (sched t) = service_in j (sched' t)

subgoal 2 (ID 511) is:
 service_in j (sched t) = service_in j (sched' t)

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


    - by rewrite !TT !service_in_implies_scheduled_in //; apply negbTE.

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

1 subgoal (ID 511)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  j : Job
  NOT_IN_NEW : ~~ scheduled_in j (sched' t')
  NOT_IN_OLD : ~~ scheduled_in j (sched t')
  t : instant
  TT : t <> t'
  ============================
  service_in j (sched t) = service_in j (sched' t)

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


    - rewrite rest_of_schedule_invariant//.

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

No more subgoals.

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


  Qed.

Based on the previous observation, we can trivially lift the invariant that jobs not involved in the replacement receive equal service to the cumulative service received in any interval.
  Corollary service_during_of_others_invariant (j: Job):
    ~~ scheduled_in j (sched' t')
    ~~ scheduled_in j (sched t')
     t1 t2,
      service_during sched j t1 t2 = service_during sched' j t1 t2.

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

1 subgoal (ID 431)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  j : Job
  ============================
  ~~ scheduled_in j (sched' t') ->
  ~~ scheduled_in j (sched t') ->
  forall t1 t2 : instant,
  service_during sched j t1 t2 = service_during sched' j t1 t2

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


  Proof.
    moveNOT_IN_NEW NOT_IN_OLD t1 t2.

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

1 subgoal (ID 435)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  j : Job
  NOT_IN_NEW : ~~ scheduled_in j (sched' t')
  NOT_IN_OLD : ~~ scheduled_in j (sched t')
  t1, t2 : instant
  ============================
  service_during sched j t1 t2 = service_during sched' j t1 t2

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


    rewrite /service_during.

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

1 subgoal (ID 442)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  j : Job
  NOT_IN_NEW : ~~ scheduled_in j (sched' t')
  NOT_IN_OLD : ~~ scheduled_in j (sched t')
  t1, t2 : instant
  ============================
  \sum_(t1 <= t < t2) service_at sched j t =
  \sum_(t1 <= t < t2) service_at sched' j t

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


    apply eq_big_natt /andP [le_t1t lt_tt2].

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

1 subgoal (ID 484)
  
  Job : JobType
  PState : Type
  H : ProcessorState Job PState
  sched : schedule PState
  t' : instant
  nstate : PState
  sched' := replace_at sched t' nstate : instant -> PState
  j : Job
  NOT_IN_NEW : ~~ scheduled_in j (sched' t')
  NOT_IN_OLD : ~~ scheduled_in j (sched t')
  t1, t2 : instant
  t : nat
  le_t1t : t1 <= t
  lt_tt2 : t < t2
  ============================
  service_at sched j t = service_at sched' j t

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


    rewrite service_at_of_others_invariant //.

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

No more subgoals.

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


  Qed.

End ReplaceAtFacts.