Skip to content
Snippets Groups Projects
Commit 4ee47809 authored by Sergey Bozhko's avatar Sergey Bozhko :eyes:
Browse files

clean up util/minmax.v

parent f193ae05
No related branches found
No related tags found
No related merge requests found
...@@ -194,7 +194,7 @@ Module RTAforEDFwithBoundedNonpreemptiveSegmentsWithArrivalCurves. ...@@ -194,7 +194,7 @@ Module RTAforEDFwithBoundedNonpreemptiveSegmentsWithArrivalCurves.
{ apply /bigmax_leq_seqP. { apply /bigmax_leq_seqP.
intros j' JINB NOTHEP. intros j' JINB NOTHEP.
apply leq_bigmax_cond_seq with apply leq_bigmax_cond_seq with
(i0 := (job_task j')) (F := fun tsk => task_max_nps tsk - 1). (x := (job_task j')) (F := fun tsk => task_max_nps tsk - 1).
{ apply H_all_jobs_from_taskset. { apply H_all_jobs_from_taskset.
apply mem_bigcat_nat_exists in JINB. apply mem_bigcat_nat_exists in JINB.
by inversion JINB as [ta' [JIN' _]]; exists ta'. } by inversion JINB as [ta' [JIN' _]]; exists ta'. }
......
...@@ -179,7 +179,7 @@ Module RTAforFPwithBoundedNonpreemptiveSegmentsWithArrivalCurves. ...@@ -179,7 +179,7 @@ Module RTAforFPwithBoundedNonpreemptiveSegmentsWithArrivalCurves.
{ apply /bigmax_leq_seqP. { apply /bigmax_leq_seqP.
intros j' JINB NOTHEP. intros j' JINB NOTHEP.
apply leq_bigmax_cond_seq with apply leq_bigmax_cond_seq with
(i0 := (job_task j')) (F := fun tsk => task_max_nps tsk - 1); last by done. (x := (job_task j')) (F := fun tsk => task_max_nps tsk - 1); last by done.
apply H_all_jobs_from_taskset. apply H_all_jobs_from_taskset.
apply mem_bigcat_nat_exists in JINB. apply mem_bigcat_nat_exists in JINB.
by inversion JINB as [ta' [JIN' _]]; exists ta'. by inversion JINB as [ta' [JIN' _]]; exists ta'.
...@@ -289,4 +289,4 @@ Module RTAforFPwithBoundedNonpreemptiveSegmentsWithArrivalCurves. ...@@ -289,4 +289,4 @@ Module RTAforFPwithBoundedNonpreemptiveSegmentsWithArrivalCurves.
End Analysis. End Analysis.
End RTAforFPwithBoundedNonpreemptiveSegmentsWithArrivalCurves. End RTAforFPwithBoundedNonpreemptiveSegmentsWithArrivalCurves.
\ No newline at end of file
...@@ -175,7 +175,7 @@ Section RTAforEDFwithBoundedNonpreemptiveSegmentsWithArrivalCurves. ...@@ -175,7 +175,7 @@ Section RTAforEDFwithBoundedNonpreemptiveSegmentsWithArrivalCurves.
by apply H_valid_model_with_bounded_nonpreemptive_segments. by apply H_valid_model_with_bounded_nonpreemptive_segments.
- apply /bigmax_leq_seqP. - apply /bigmax_leq_seqP.
intros j' JINB NOTHEP. intros j' JINB NOTHEP.
apply leq_bigmax_cond_seq with (i0 := (job_task j')) (F := fun tsk => task_max_nonpreemptive_segment tsk - 1). apply leq_bigmax_cond_seq with (x := (job_task j')) (F := fun tsk => task_max_nonpreemptive_segment tsk - 1).
{ apply H_all_jobs_from_taskset. { apply H_all_jobs_from_taskset.
apply mem_bigcat_nat_exists in JINB. apply mem_bigcat_nat_exists in JINB.
by inversion JINB as [ta' [JIN' _]]; exists ta'. } by inversion JINB as [ta' [JIN' _]]; exists ta'. }
......
...@@ -153,7 +153,7 @@ Section RTAforFPwithBoundedNonpreemptiveSegmentsWithArrivalCurves. ...@@ -153,7 +153,7 @@ Section RTAforFPwithBoundedNonpreemptiveSegmentsWithArrivalCurves.
{ apply /bigmax_leq_seqP. { apply /bigmax_leq_seqP.
intros j' JINB NOTHEP. intros j' JINB NOTHEP.
apply leq_bigmax_cond_seq with apply leq_bigmax_cond_seq with
(i0 := (job_task j')) (F := fun tsk => task_max_nonpreemptive_segment tsk - 1); last by done. (x := (job_task j')) (F := fun tsk => task_max_nonpreemptive_segment tsk - 1); last by done.
apply H_all_jobs_from_taskset. apply H_all_jobs_from_taskset.
apply mem_bigcat_nat_exists in JINB. apply mem_bigcat_nat_exists in JINB.
by inversion JINB as [ta' [JIN' _]]; exists ta'. by inversion JINB as [ta' [JIN' _]]; exists ta'.
......
Require Export prosa.util.notation prosa.util.nat prosa.util.list. Require Export prosa.util.notation prosa.util.nat prosa.util.list.
From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq fintype bigop. From mathcomp Require Import ssreflect ssrbool eqtype ssrnat seq fintype bigop.
(* Additional lemmas about max. *) (** Additional lemmas about [BigMax]. *)
Section ExtraLemmas. Section ExtraLemmas.
Lemma leq_bigmax_cond_seq (T : eqType) (P : pred T) (r : seq T) F i0 : (** Given a function [F], a predicate [P], and a sequence [xs], we
i0 \in r -> P i0 -> F i0 <= \max_(i <- r | P i) F i. show that [F x <= max { F i | ∀ i ∈ xs, P i}] for any [x] in
Proof. [xs]. *)
intros IN0 Pi0; by rewrite (big_rem i0) //= Pi0 leq_maxl. Lemma leq_bigmax_cond_seq :
Qed. forall {X : eqType} (F : X -> nat) (P : pred X) (xs : seq X) (x : X),
x \in xs -> P x -> F x <= \max_(i <- xs | P i) F i.
Lemma bigmax_sup_seq: Proof. by intros * IN Px; rewrite (big_rem x) //= Px leq_maxl. Qed.
forall (T: eqType) (i: T) r (P: pred T) m F,
i \in r -> P i -> m <= F i -> m <= \max_(i <- r| P i) F i.
Proof.
intros.
induction r.
- by rewrite in_nil in H.
move: H; rewrite in_cons; move => /orP [/eqP EQA | IN].
{
clear IHr; subst a.
rewrite (big_rem i) //=; last by rewrite in_cons; apply/orP; left.
apply leq_trans with (F i); first by done.
by rewrite H0 leq_maxl.
}
{
apply leq_trans with (\max_(i0 <- r | P i0) F i0); first by apply IHr.
rewrite [in X in _ <= X](big_rem a) //=; last by rewrite in_cons; apply/orP; left.
have Ob: a == a; first by done.
by rewrite Ob leq_maxr.
}
Qed.
Lemma bigmax_leq_seqP (T : eqType) (P : pred T) (r : seq T) F m : (** Next, we show that the fact [max { F i | ∀ i ∈ xs, P i} <= m] for
reflect (forall i, i \in r -> P i -> F i <= m) some [m] is equivalent to the fact that [∀ x ∈ xs, P x -> F x <= m]. *)
(\max_(i <- r | P i) F i <= m). Lemma bigmax_leq_seqP :
forall {X : eqType} (F : X -> nat) (P : pred X) (xs : seq X) (m : nat),
reflect (forall x, x \in xs -> P x -> F x <= m)
(\max_(x <- xs | P x) F x <= m).
Proof. Proof.
apply: (iffP idP) => leFm => [i IINR Pi|]; intros *.
first by apply: leq_trans leFm; apply leq_bigmax_cond_seq. apply: (iffP idP) => leFm => [i IINR Pi|].
rewrite big_seq_cond; elim/big_ind: _ => // m1 m2. - by apply: leq_trans leFm; apply leq_bigmax_cond_seq.
by intros; rewrite geq_max; apply/andP; split. - rewrite big_seq_cond; elim/big_ind: _ => // m1 m2.
by move: m2 => /andP [M1IN Pm1]; apply: leFm. + by intros; rewrite geq_max; apply/andP; split.
+ by move: m2 => /andP [M1IN Pm1]; apply: leFm.
Qed. Qed.
Lemma leq_big_max (T : eqType) (P : pred T) (r : seq T) F1 F2 : (** Given two functions [F1] and [F2], a predicate [P], and sequence
(forall i, i \in r -> P i -> F1 i <= F2 i) -> [xs], we show that if [F1 x <= F2 x] for any [x \in xs] such that
\max_(i <- r | P i) F1 i <= \max_(i <- r | P i) F2 i. [P x], then [max] of [{F1 x | ∀ x ∈ xs, P x}] is bounded by the
[max] of [{F2 x | ∀ x ∈ xs, P x}]. *)
Lemma leq_big_max :
forall {X : eqType} (F1 F2 : X -> nat) (P : pred X) (xs : seq X),
(forall x, x \in xs -> P x -> F1 x <= F2 x) ->
\max_(x <- xs | P x) F1 x <= \max_(x <- xs | P x) F2 x.
Proof. Proof.
intros; apply /bigmax_leq_seqP; intros. intros * ALL; apply /bigmax_leq_seqP; intros * IN Px.
specialize (H i); feed_n 2 H; try(done). specialize (ALL x); feed_n 2 ALL; try done.
rewrite (big_rem i) //=; rewrite H1. rewrite (big_rem x) //=; rewrite Px.
by apply leq_trans with (F2 i); [ | rewrite leq_maxl]. by apply leq_trans with (F2 x); [ | rewrite leq_maxl].
Qed. Qed.
Lemma bigmax_ord_ltn_identity n : (** We show that for a positive [n], [max] of [{0, …, n-1}] is
n > 0 -> smaller than [n]. *)
\max_(i < n) i < n. Lemma bigmax_ord_ltn_identity :
forall (n : nat),
n > 0 ->
\max_(i < n) i < n.
Proof. Proof.
intros LT. intros [ | n] POS; first by rewrite ltn0 in POS.
destruct n; first by rewrite ltn0 in LT. clear POS; induction n; first by rewrite big_ord_recr /= big_ord0 maxn0.
clear LT.
induction n; first by rewrite big_ord_recr /= big_ord0 maxn0.
rewrite big_ord_recr /=. rewrite big_ord_recr /=.
by unfold maxn at 1; rewrite IHn. by rewrite /maxn IHn.
Qed. Qed.
Lemma bigmax_ltn_ord n (P : pred nat) (i0: 'I_n) : (** We state the next lemma in terms of _ordinals_. Given a natural
number [n], a predicate [P], and an ordinal [i0 ∈ {0, …, n-1}]
satisfying [P], we show that [max {i | P i} < n]. Note that the
element satisfying [P] is needed to prove that [{i | P i}] is
not empty. *)
Lemma bigmax_ltn_ord :
forall (n : nat) (P : pred nat) (i0 : 'I_n),
P i0 -> P i0 ->
\max_(i < n | P i) i < n. \max_(i < n | P i) i < n.
Proof. Proof.
intros LT. intros n P i0 Pi.
destruct n; first by destruct i0 as [i0 P0]; move: (P0) => P0'; rewrite ltn0 in P0'. destruct n; first by destruct i0 as [i0 P0]; move: (P0) => P0'; rewrite ltn0 in P0'.
rewrite big_mkcond. rewrite big_mkcond.
apply leq_ltn_trans with (n := \max_(i < n.+1) i). apply leq_ltn_trans with (n := \max_(i < n.+1) i).
{ - apply/bigmax_leqP; ins.
apply/bigmax_leqP; ins.
destruct (P i); last by done. destruct (P i); last by done.
by apply leq_bigmax_cond. by apply leq_bigmax_cond.
} - by apply bigmax_ord_ltn_identity.
by apply bigmax_ord_ltn_identity.
Qed. Qed.
Lemma bigmax_pred n (P : pred nat) (i0: 'I_n) : (** Finally, we show that, given a natural number [n], a predicate
P (i0) -> [P], and an ordinal [i0 ∈ {0, …, n-1}] satisfying [P],
P (\max_(i < n | P i) i). [max {i | P i} < n] also satisfies [P]. *)
Lemma bigmax_pred :
forall (n : nat) (P : pred nat) (i0 : 'I_n),
P i0 ->
P (\max_(i < n | P i) i).
Proof. Proof.
intros PRED. intros * Pi0.
induction n. induction n.
{ { by destruct i0 as [i0 P0]; move: (P0) => P1; rewrite ltn0 in P1. }
destruct i0 as [i0 P0]. { rewrite big_mkcond big_ord_recr /=.
by move: (P0) => P1; rewrite ltn0 in P1. destruct (P n) eqn:Pn.
} { destruct n; first by rewrite big_ord0 maxn0.
rewrite big_mkcond big_ord_recr /=. destruct (P n) eqn:Pn. unfold maxn at 1.
{ destruct (\max_(i < n.+1) (match P (@nat_of_ord (S n) i) return nat with
destruct n; first by rewrite big_ord0 maxn0. | true => @nat_of_ord (S n) i
unfold maxn at 1. | false => O
destruct (\max_(i < n.+1) (match P (@nat_of_ord (S n) i) return nat with end) < n.+1) eqn:Pi; first by rewrite Pi.
| true => @nat_of_ord (S n) i exfalso.
| false => O apply negbT in Pi; move: Pi => /negP BUG; apply: BUG.
end) < n.+1) eqn:Pi; first by rewrite Pi. apply leq_ltn_trans with (n := \max_(i < n.+1) i).
exfalso. { apply/bigmax_leqP; rewrite //= => i _.
apply negbT in Pi; move: Pi => /negP BUG. by destruct (P i); first apply leq_bigmax_cond.
apply BUG. }
apply leq_ltn_trans with (n := \max_(i < n.+1) i). { by apply bigmax_ord_ltn_identity. }
{
apply/bigmax_leqP; ins.
destruct (P i); last by done.
by apply leq_bigmax_cond.
} }
by apply bigmax_ord_ltn_identity. { rewrite maxn0 -big_mkcond /=.
} have LT: i0 < n; last by rewrite (IHn (Ordinal LT)).
{
rewrite maxn0.
rewrite -big_mkcond /=.
have LT: i0 < n.
{
rewrite ltn_neqAle; apply/andP; split; rewrite ltn_neqAle; apply/andP; split;
last by rewrite -ltnS; apply ltn_ord. last by rewrite -ltnS; apply ltn_ord.
apply/negP; move => /eqP BUG. apply/negP; move => /eqP BUG.
by rewrite -BUG PRED in Pn. by rewrite -BUG Pi0 in Pn.
} }
by rewrite (IHn (Ordinal LT)).
} }
Qed. Qed.
End ExtraLemmas. End ExtraLemmas.
\ No newline at end of file
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment