Library probsa.probability.dominance_relation

(* --------------------------------- ProBsa ---------------------------------- *)
From probsa.util Require Export etime.
From probsa.probability Require Export nrvar cdf.

(* ------------------------------- Definition -------------------------------- *)
We introduce a generic notion of dominance between elements of two arbitrary types. The expression a b denotes the fact that a : A is dominated by b : B.
Class DominanceRelation (A B : Type) :=
  dominates : A B Prop.
Notation "X ⪯ Y" := (dominates X Y) (at level 55) : probability_scope.

(* -------------------------------- Instances -------------------------------- *)

Instances

nat R vs nat R

Given two functions f and g, we say that f g iff the graph of f is always above the g's graph. That is, x, f(x) g(x).
Definition le_func (f g : nat R) :=
   n, f n g n.

Instance dominance_func : DominanceRelation (nat R) (nat R) :=
  { dominates := le_func }.

Lemma func_dom_refl :
   (f : nat R), f f.

Lemma func_dom_trans :
   (f g h : nat R),
    f g g h f h.

nrvar vs nrvar

Given two random variables X and Y, we say that X Y iff X's CDF is always above Y's CDF. That is, n, 𝔽[X](n) 𝔽[Y](n).
Definition le_nrvar { } {μX : measure } {μY : measure } (X : nrvar μX) (Y : nrvar μY) :=
   𝔽<μX>{[ X ]} 𝔽<μY>{[ Y ]}.

Instance dominance_nrvar :
   {ΩX ΩY} {μX : measure ΩX} {μY : measure ΩY},
    DominanceRelation (nrvar μX) (nrvar μY) :=
  { dominates := le_nrvar }.

Lemma nrvar_dom_refl :
   {Ω} {μ : measure Ω} (X : nrvar μ),
    X X.

Lemma nrvar_dom_trans :
   {} {μX : measure } (X : nrvar μX),
   {} {μY : measure } (Y : nrvar μY),
   {} {μZ : measure } (Z : nrvar μZ),
    X Y Y Z X Z.

nrvar vs nat R

Similarly, when a random variable X is compared with a function f : [0,1], we say that X f iff X's CDF is always above f. That is, x, 𝔽[X](x) f x.
Definition le_nrvar_func {Ω} {μ : measure Ω} (X : nrvar μ) (f : nat R) :=
  𝔽<μ>{[ X ]} f.

Instance dominance_nrvar_funct :
   {Ω} {μ : measure Ω}, DominanceRelation (nrvar μ) (nat R) :=
  { dominates := le_nrvar_func }.

Instance nat_etimervar_pred_ltop :
   {Ω} {μ : measure Ω}, LtOp nat (rvar μ [eqType of etime]) (pred Ω) :=
  { lt_op t X := fun ωexceeds (X ω) t }.

Definition le_etime_rvar { } {μX : measure } {μY : measure }
           (X : rvar μX [eqType of etime]) (Y : rvar μY [eqType of etime]) :=
   (t : nat), <μX>{[ t ⟨<⟩ X ]} <μY>{[ t ⟨<⟩ Y ]}.

Instance dominance_etime_rvar :
   {ΩX ΩY} {μX : measure ΩX} {μY : measure ΩY},
    DominanceRelation (rvar μX [eqType of etime]) (rvar μY [eqType of etime]) :=
  { dominates := le_etime_rvar }.

Lemma etime_dom_refl :
   {Ω} {μ : measure Ω} (X : rvar μ [eqType of etime]),
    X X.

Lemma etime_dom_trans :
   {} {μX : measure } (X : rvar μX [eqType of etime]),
   {} {μY : measure } (Y : rvar μY [eqType of etime]),
   {} {μZ : measure } (Z : rvar μZ [eqType of etime]),
    X Y Y Z X Z.