Library prosa.model.processor.varspeed


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

Welcome to Coq 8.11.2 (June 2020)

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


From mathcomp Require Import all_ssreflect.
Require Export prosa.behavior.all.

In the following, we define a model of processors with variable execution speeds.
NB: For now, the definition serves only to document how this can be done; it is not actually used anywhere in the library.

Section State.

Consider any type of jobs.
  Variable Job: JobType.

We define the state of a variable-speed processor at a given time to be one of two possible cases: either a specific job is scheduled and progresses with a specific speed, or the processor is idle.
  Inductive processor_state :=
    Idle
  | Progress (j : Job) (speed : nat).

Next, we define the semantics of the variable-speed processor state.
  Section Service.

Consider any job [j].
    Variable j : Job.

Job [j] is scheduled in a given state [s] if [s] is not idle and [j] matches the job recorded in [s].
    Definition varspeed_scheduled_on (s : processor_state) (_ : unit) : bool :=
      match s with
      | Idlefalse
      | Progress j' _j' == j
      end.

If it is scheduled in state [s], job [j] receives service proportional to the speed recorded in the state.
    Definition varspeed_service_in (s : processor_state) : nat :=
      match s with
      | Idle ⇒ 0
      | Progress j' speedif j' == j then speed else 0
      end.

  End Service.

Finally, we connect the above definitions to the generic Prosa interface for processor states.
  Global Program Instance pstate_instance : ProcessorState Job processor_state :=
    {
      scheduled_on := varspeed_scheduled_on;
      service_in := varspeed_service_in
    }.
  Next Obligation.

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

1 subgoal (ID 583)
  
  Job : JobType
  j : Job
  s : processor_state
  H : ~~ [exists c, varspeed_scheduled_on j s c]
  ============================
  varspeed_service_in j s = 0

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


    move: H.

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

1 subgoal (ID 590)
  
  Job : JobType
  j : Job
  s : processor_state
  ============================
  ~~ [exists c, varspeed_scheduled_on j s c] -> varspeed_service_in j s = 0

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


    case: s=>//= j' v /existsP.

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

1 subgoal (ID 669)
  
  Job : JobType
  j, j' : Job
  v : nat
  ============================
  ~ (exists _ : unit_finType, j' == j) -> (if j' == j then v else 0) = 0

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


    case: ifP=>//=_[].

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

1 subgoal (ID 750)
  
  Job : JobType
  j, j' : Job
  v : nat
  ============================
  exists _ : unit, true

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


    by .

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

No more subgoals.

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


  Defined.

End State.