From 6a1dfeb365f91db6ce48b1b123c07d22cb7011ca Mon Sep 17 00:00:00 2001
From: Kimaya Bedarkar <kbedarka@mpi-sws.org>
Date: Wed, 19 Feb 2025 10:57:51 +0100
Subject: [PATCH] add proc state facts

---
 analysis/facts/model/ideal_uni_exceed.v | 133 ++++++++++++++++++++++++
 1 file changed, 133 insertions(+)
 create mode 100644 analysis/facts/model/ideal_uni_exceed.v

diff --git a/analysis/facts/model/ideal_uni_exceed.v b/analysis/facts/model/ideal_uni_exceed.v
new file mode 100644
index 000000000..b30a00b3c
--- /dev/null
+++ b/analysis/facts/model/ideal_uni_exceed.v
@@ -0,0 +1,133 @@
+Require Import prosa.model.processor.ideal_uni_exceed.
+Require Export prosa.model.processor.platform_properties.
+
+(** In the following section, we prove some general properties about the
+    ideal uniprocessor with exceedance processor state. *)
+Section ExceedanceProcStateProperties.
+
+  (** In this section, we consider any type of jobs. *)
+  Context `{Job : JobType}.
+
+  (** First we prove that the processor state provides unit supply. *)
+  Lemma eps_is_unit_supply : unit_supply_proc_model (exceedance_proc_state Job).
+  Proof.
+    rewrite /unit_supply_proc_model.
+    move => s.
+    destruct s => //=.
+    all : by rewrite /supply_in //=;rewrite sum_unit1.
+  Qed.
+
+  (** The usually opaque predicates are made transparent for this section
+      to enable us to prove some processor properties. *)
+  Local Transparent scheduled_in scheduled_on service_on.
+
+  (** A job [j] is said to be [scheduled_at] some time [t] if and only if the
+      processor state at that time is [NominalExecution j] or
+      [ExceedanceExecution j]. *)
+   Lemma scheduled_at_procstate (sched : schedule (exceedance_proc_state Job)) j t:
+    scheduled_at sched j t <->
+      sched t = NominalExecution j \/ sched t = ExceedanceExecution j.
+  Proof.
+    split.
+    -  move => SCHED.
+      rewrite /scheduled_at /scheduled_in in SCHED.
+      move : SCHED => /existsP [c SCHEDON].
+      unfold scheduled_on in SCHEDON.
+      simpl in SCHEDON.
+      unfold exceedance_scheduled_on in *.
+      destruct (sched t); try done.
+      + left.
+        by move : SCHEDON => /eqP <-.
+      + right.
+        by move : SCHEDON => /eqP <-.
+    - rewrite /scheduled_at.
+      move => [-> | -> ]; rewrite /scheduled_in /scheduled_on //=; by apply /existsP => //=.
+  Qed.
+
+
+  (** Next we prove that the processor model under consideration is a uniprocessor
+      model i.e., only one job can be scheduled at any time instant. *)
+  Lemma eps_is_uniproc : uniprocessor_model (exceedance_proc_state Job).
+  Proof.
+    rewrite /uniprocessor_model.
+    move => j1 j2 sched t SCHED1 SCHED2.
+    rewrite /scheduled_at /scheduled_in in SCHED1, SCHED2.
+    move : SCHED1 => /existsP [c SCHED1].
+    move : SCHED2 => /existsP [c' SCHED2].
+    rewrite  /scheduled_on //= /exceedance_scheduled_on in SCHED1, SCHED2.
+    destruct (sched t) eqn: EQ; try done.
+    - move : SCHED1 => /eqP <-.
+      by move : SCHED2 => /eqP <-.
+    - move : SCHED1 => /eqP <-.
+      by move : SCHED2 => /eqP <-.
+  Qed.
+
+  (** Next, we prove that the processor model under consideration is fully
+      consuming i.e., all the supply offered is consumed by the scheduled job. *)
+  Lemma eps_is_fully_consuming :
+    fully_consuming_proc_model (exceedance_proc_state Job).
+  Proof.
+    rewrite /fully_consuming_proc_model.
+    move => j sched t SCHED1.
+    rewrite /scheduled_at /scheduled_in in SCHED1.
+    apply scheduled_at_procstate in SCHED1.
+    move : SCHED1 => [SCHED1 | SCHED1].
+    - by rewrite /service_at /service_in /exceedance_service_on /supply_at /supply_in
+      !/exceedance_service_on //= SCHED1 sum_unit1 // sum_unit1 //
+      /exceedance_service_on /exceedance_supply_on eq_refl.
+    - by rewrite /service_at /service_in /exceedance_service_on /supply_at /supply_in
+        !/exceedance_service_on //= SCHED1 sum_unit1 // sum_unit1 //
+        /exceedance_service_on /exceedance_supply_on eq_refl.
+  Qed.
+
+  (** Finally we prove that the processor model under consideration is unit
+      service i.e., it provides at most one unit of service at any time instant. *)
+  Lemma eps_is_unit_service : unit_service_proc_model (exceedance_proc_state Job).
+  Proof.
+    rewrite /unit_service_proc_model.
+    move => j pstate.
+    rewrite /service_in /service_on //= sum_unit1 /exceedance_service_on.
+    destruct pstate =>//=.
+    by apply leq_b1.
+  Qed.
+
+End ExceedanceProcStateProperties.
+
+(** Next, we prove some general properties specific to the
+    ideal uniprocessor with exceedance processor state. *)
+Section ExceedanceProcStateProperties1.
+
+  (** In this section we consider any type of jobs. *)
+  Context `{Job : JobType}.
+
+  (** First let us define the notion of exceedance execution. A processor state
+      is said to be an exceedance execution state if is equal to
+      [ExceedanceExecution j] for some [j]. *)
+  Definition is_exceedance_exec (pstate : (exceedance_proc_state Job)) :=
+    match pstate with
+    |ExceedanceExecution _ => true
+    |_ => false
+    end.
+  Local Transparent scheduled_in scheduled_on service_on.
+
+
+  (** Next, let us consider any schedule of the [exceedance_proc_state]. *)
+  Variable sched : schedule (exceedance_proc_state Job).
+
+  (** We prove that the schedule being in blackout at any time [t] is equivalent
+      to the processor state at [t] being an exceedance execution state. *)
+  Lemma blackout_implies_exceedance_execution t :
+    is_blackout sched t = is_exceedance_exec (sched t).
+  Proof.
+    rewrite /is_blackout /has_supply /supply_at /supply_in
+    /is_exceedance_exec //= sum_unit1 /exceedance_supply_on.
+    destruct (sched t) eqn : PSTATE; try done.
+  Qed.
+
+End ExceedanceProcStateProperties1.
+
+Global Hint Resolve
+  eps_is_uniproc
+  eps_is_unit_supply
+  eps_is_fully_consuming
+  : basic_rt_facts.
-- 
GitLab